src/HOL/Library/Formal_Power_Series.thy
author wenzelm
Sun, 18 Aug 2013 19:59:19 +0200
changeset 53077 a1b3784f8129
parent 52903 6c89225ddeba
child 53195 e4b18828a817
permissions -rw-r--r--
more symbols;
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
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     5
header{* A formalization of formal power series *}
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
51542
738598beeb26 tuned imports;
wenzelm
parents: 51489
diff changeset
     8
imports Binomial
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
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
    11
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
    12
subsection {* The type of formal power series*}
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
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
    29
text{* Definition of the basic elements 0 and 1 and the basic operations of addition,
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
    30
  negation and multiplication *}
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
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    34
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    35
definition fps_zero_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    36
  "0 = Abs_fps (\<lambda>n. 0)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    37
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    38
instance ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    39
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    40
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    41
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
    42
  unfolding fps_zero_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    43
36409
d323e7773aa8 use new classes (linordered_)field_inverse_zero
haftmann
parents: 36350
diff changeset
    44
instantiation fps :: ("{one, zero}") one
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    45
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    46
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    47
definition fps_one_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    48
  "1 = Abs_fps (\<lambda>n. if n = 0 then 1 else 0)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    49
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    50
instance ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    51
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    52
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
    53
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
    54
  unfolding fps_one_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    55
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    56
instantiation fps :: (plus)  plus
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    57
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    58
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    59
definition fps_plus_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    60
  "op + = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n + g $ n))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    61
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    62
instance ..
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_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
    66
  unfolding fps_plus_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 :: (minus) minus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    69
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    70
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    71
definition fps_minus_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    72
  "op - = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n - g $ n))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    73
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    74
instance ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    75
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    76
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    77
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
    78
  unfolding fps_minus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    79
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    80
instantiation fps :: (uminus) uminus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    81
begin
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
definition fps_uminus_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    84
  "uminus = (\<lambda>f. Abs_fps (\<lambda>n. - (f $ n)))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    85
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    86
instance ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    87
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    88
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    89
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
    90
  unfolding fps_uminus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    91
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    92
instantiation fps :: ("{comm_monoid_add, times}")  times
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    93
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    94
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    95
definition fps_times_def:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    96
  "op * = (\<lambda>f g. Abs_fps (\<lambda>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
    97
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    98
instance ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    99
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   100
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   101
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
   102
  unfolding fps_times_def by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   103
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   104
declare atLeastAtMost_iff [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   105
declare Bex_def [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   106
declare Ball_def [presburger]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   107
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   108
lemma mult_delta_left:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   109
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   110
  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
   111
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   112
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   113
lemma mult_delta_right:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   114
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   115
  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
   116
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   117
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   118
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
   119
  by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   120
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   121
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
   122
  by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   123
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   124
subsection{* Formal power series form a commutative ring with unity, if the range of sequences
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   125
  they represent is a commutative ring with unity*}
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
instance fps :: (semigroup_add) semigroup_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   128
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   129
  fix a b c :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   130
  show "a + b + c = a + (b + c)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   131
    by (simp add: fps_ext add_assoc)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   132
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   133
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   134
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
   135
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   136
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   137
  show "a + b = b + a"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   138
    by (simp add: fps_ext add_commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   139
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   140
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   141
lemma fps_mult_assoc_lemma:
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   142
  fixes k :: nat and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   143
  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
   144
         (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   145
  by (induct k) (simp_all add: Suc_diff_le setsum_addf add_assoc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   146
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   147
instance fps :: (semiring_0) semigroup_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   148
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   149
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   150
  show "(a * b) * c = a * (b * c)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   151
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   152
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   153
    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
   154
          (\<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
   155
      by (rule fps_mult_assoc_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   156
    then show "((a * b) * c) $ n = (a * (b * c)) $ n"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   157
      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
   158
  qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   159
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   160
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   161
lemma fps_mult_commute_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   162
  fixes n :: nat
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   163
    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
   164
  shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   165
proof (rule setsum_reindex_cong)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   166
  show "inj_on (\<lambda>i. n - i) {0..n}"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   167
    by (rule inj_onI) simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   168
  show "{0..n} = (\<lambda>i. n - i) ` {0..n}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   169
    apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   170
    apply (rule_tac x = "n - x" in image_eqI)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   171
    apply simp_all
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   172
    done
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   173
next
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   174
  fix i
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   175
  assume "i \<in> {0..n}"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   176
  then have "n - (n - i) = i" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   177
  then show "f (n - i) i = f (n - i) (n - (n - i))" by simp
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   178
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   179
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   180
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
   181
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   182
  fix a b :: "'a fps"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   183
  show "a * b = b * a"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   184
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   185
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   186
    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
   187
      by (rule fps_mult_commute_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   188
    then show "(a * b) $ n = (b * a) $ n"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   189
      by (simp add: fps_mult_nth mult_commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   190
  qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   191
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   192
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   193
instance fps :: (monoid_add) monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   194
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   195
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   196
  show "0 + a = a" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   197
  show "a + 0 = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   198
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   199
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   200
instance fps :: (comm_monoid_add) comm_monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   201
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   202
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   203
  show "0 + a = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   204
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   205
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   206
instance fps :: (semiring_1) monoid_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   207
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   208
  fix a :: "'a fps"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   209
  show "1 * a = a" by (simp add: fps_ext fps_mult_nth mult_delta_left setsum_delta)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   210
  show "a * 1 = a" 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
   211
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   212
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   213
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
   214
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   215
  fix a b c :: "'a fps"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   216
  { assume "a + b = a + c" then show "b = c" by (simp add: expand_fps_eq) }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   217
  { assume "b + a = c + a" then show "b = c" by (simp add: expand_fps_eq) }
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   218
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   219
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   220
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
   221
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   222
  fix a b c :: "'a fps"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   223
  assume "a + b = a + c"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   224
  then show "b = c" by (simp add: expand_fps_eq)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   225
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   226
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   227
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
   228
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   229
instance fps :: (group_add) group_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   230
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   231
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   232
  show "- a + a = 0" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   233
  show "a - b = a + - b" by (simp add: fps_ext diff_minus)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   234
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   235
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   236
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
   237
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   238
  fix a b :: "'a fps"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   239
  show "- a + a = 0" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   240
  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
   241
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   242
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   243
instance fps :: (zero_neq_one) zero_neq_one
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   244
  by default (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   245
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   246
instance fps :: (semiring_0) semiring
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   247
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   248
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   249
  show "(a + b) * c = a * c + b * c"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
   250
    by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum_addf)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   251
  show "a * (b + c) = a * b + a * c"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
   252
    by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum_addf)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   253
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   254
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   255
instance fps :: (semiring_0) semiring_0
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   256
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   257
  fix a:: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   258
  show "0 * a = 0" by (simp add: fps_ext fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   259
  show "a * 0 = 0" by (simp add: fps_ext fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   260
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   261
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   262
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
   263
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
   264
subsection {* Selection of the nth power of the implicit variable in the infinite sum*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   265
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   266
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
   267
  by (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   268
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   269
lemma fps_nonzero_nth_minimal: "f \<noteq> 0 \<longleftrightarrow> (\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m < n. f $ m = 0))"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   270
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   271
  let ?n = "LEAST n. f $ n \<noteq> 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   272
  assume "f \<noteq> 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   273
  then have "\<exists>n. f $ n \<noteq> 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   274
    by (simp add: fps_nonzero_nth)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   275
  then have "f $ ?n \<noteq> 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   276
    by (rule LeastI_ex)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   277
  moreover have "\<forall>m<?n. f $ m = 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   278
    by (auto dest: not_less_Least)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   279
  ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   280
  then show "\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m<n. f $ m = 0)" ..
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   281
next
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   282
  assume "\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m<n. f $ m = 0)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   283
  then show "f \<noteq> 0" by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   284
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   285
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   286
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
   287
  by (rule expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   288
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   289
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
   290
proof (cases "finite S")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   291
  case True
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   292
  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
   293
next
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   294
  case False
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   295
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   296
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   297
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
   298
subsection{* Injection of the basic ring elements and multiplication by scalars *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   299
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   300
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
   301
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   302
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
   303
  unfolding fps_const_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   304
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   305
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
   306
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   307
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   308
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
   309
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   310
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   311
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
   312
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   313
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   314
lemma fps_const_add [simp]: "fps_const (c::'a\<Colon>monoid_add) + fps_const d = fps_const (c + d)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   315
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   316
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
   317
lemma fps_const_sub [simp]: "fps_const (c::'a\<Colon>group_add) - fps_const d = fps_const (c - d)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
   318
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   319
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   320
lemma fps_const_mult[simp]: "fps_const (c::'a\<Colon>ring) * 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
   321
  by (simp add: fps_eq_iff fps_mult_nth setsum_0')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   322
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   323
lemma fps_const_add_left: "fps_const (c::'a\<Colon>monoid_add) + f =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   324
    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
   325
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   326
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   327
lemma fps_const_add_right: "f + fps_const (c::'a\<Colon>monoid_add) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   328
    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
   329
  by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   330
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   331
lemma fps_const_mult_left: "fps_const (c::'a\<Colon>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
   332
  unfolding fps_eq_iff fps_mult_nth
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   333
  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
   334
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   335
lemma fps_const_mult_right: "f * fps_const (c::'a\<Colon>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
   336
  unfolding fps_eq_iff fps_mult_nth
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   337
  by (simp add: fps_const_def mult_delta_right setsum_delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   338
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   339
lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   340
  by (simp add: fps_mult_nth mult_delta_left setsum_delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   341
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   342
lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   343
  by (simp add: fps_mult_nth mult_delta_right setsum_delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   344
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
   345
subsection {* Formal power series form an integral domain*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   346
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   347
instance fps :: (ring) ring ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   348
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   349
instance fps :: (ring_1) ring_1
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
   350
  by (intro_classes, auto simp add: diff_minus distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   351
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   352
instance fps :: (comm_ring_1) comm_ring_1
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
   353
  by (intro_classes, auto simp add: diff_minus distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   354
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   355
instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   356
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   357
  fix a b :: "'a fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   358
  assume a0: "a \<noteq> 0" and b0: "b \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   359
  then obtain i j where i: "a$i\<noteq>0" "\<forall>k<i. a$k=0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   360
    and j: "b$j \<noteq>0" "\<forall>k<j. b$k =0" unfolding fps_nonzero_nth_minimal
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   361
    by blast+
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   362
  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
   363
    by (rule fps_mult_nth)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   364
  also have "\<dots> = (a$i * b$(i+j-i)) + (\<Sum>k\<in>{0..i+j}-{i}. a$k * b$(i+j-k))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   365
    by (rule setsum_diff1') simp_all
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   366
  also have "(\<Sum>k\<in>{0..i+j}-{i}. a$k * b$(i+j-k)) = 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   367
    proof (rule setsum_0' [rule_format])
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   368
      fix k assume "k \<in> {0..i+j} - {i}"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   369
      then have "k < i \<or> i+j-k < j" by auto
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   370
      then show "a$k * b$(i+j-k) = 0" using i j by auto
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   371
    qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   372
  also have "a$i * b$(i+j-i) + 0 = a$i * b$j" by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   373
  also have "a$i * b$j \<noteq> 0" using i j by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   374
  finally have "(a*b) $ (i+j) \<noteq> 0" .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   375
  then show "a*b \<noteq> 0" unfolding fps_nonzero_nth by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   376
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   377
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   378
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
   379
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   380
instance fps :: (idom) idom ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   381
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   382
lemma numeral_fps_const: "numeral k = fps_const (numeral k)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   383
  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
   384
    fps_const_add [symmetric])
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   385
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   386
lemma neg_numeral_fps_const: "neg_numeral k = fps_const (neg_numeral k)"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   387
  by (simp only: neg_numeral_def numeral_fps_const fps_const_neg)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   388
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   389
subsection{* The eXtractor series X*}
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   390
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   391
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
   392
  by (induct n) auto
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   393
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   394
definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   395
lemma X_mult_nth[simp]: "(X * (f :: ('a::semiring_1) fps)) $n = (if n = 0 then 0 else f $ (n - 1))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   396
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   397
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   398
    assume n: "n \<noteq> 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   399
    have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   400
      by (simp add: fps_mult_nth)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   401
    also have "\<dots> = f $ (n - 1)"
46757
ad878aff9c15 removing finiteness goals
bulwahn
parents: 46131
diff changeset
   402
      using n by (simp add: X_def mult_delta_left setsum_delta)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   403
    finally have ?thesis using n by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   404
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   405
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   406
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   407
    assume n: "n=0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   408
    hence ?thesis by (simp add: fps_mult_nth X_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   409
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   410
  ultimately show ?thesis by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   411
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   412
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   413
lemma X_mult_right_nth[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   414
    "((f :: ('a::comm_semiring_1) fps) * X) $n = (if n = 0 then 0 else f $ (n - 1))"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   415
  by (metis X_mult_nth mult_commute)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   416
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   417
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
   418
proof (induct k)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   419
  case 0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   420
  thus ?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
   421
next
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   422
  case (Suc k)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   423
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   424
    fix m
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   425
    have "(X^Suc k) $ m = (if m = 0 then (0::'a) else (X^k) $ (m - 1))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   426
      by (simp del: One_nat_def)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   427
    then have "(X^Suc k) $ m = (if m = Suc k then (1::'a) else 0)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   428
      using Suc.hyps by (auto cong del: if_weak_cong)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   429
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   430
  then show ?case by (simp add: fps_eq_iff)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   431
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   432
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   433
lemma X_power_mult_nth:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   434
    "(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
   435
  apply (induct k arbitrary: n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   436
  apply simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   437
  unfolding power_Suc mult_assoc
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   438
  apply (case_tac n)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   439
  apply auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   440
  done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   441
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   442
lemma X_power_mult_right_nth:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   443
    "((f :: ('a::comm_ring_1) fps) * X^k) $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
   444
  by (metis X_power_mult_nth mult_commute)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   445
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   446
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   447
subsection{* Formal Power series form a metric space *}
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   448
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   449
definition (in dist) "ball x r = {y. dist y x < r}"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   450
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   451
instantiation fps :: (comm_ring_1) dist
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   452
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   453
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   454
definition
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   455
  dist_fps_def: "dist (a::'a fps) b =
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   456
    (if (\<exists>n. a$n \<noteq> b$n) then inverse (2 ^ The (leastP (\<lambda>n. a$n \<noteq> b$n))) else 0)"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   457
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   458
lemma dist_fps_ge0: "dist (a::'a fps) b \<ge> 0"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   459
  by (simp add: dist_fps_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   460
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   461
lemma dist_fps_sym: "dist (a::'a fps) b = dist b a"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   462
  apply (auto simp add: dist_fps_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   463
  apply (rule cong[OF refl, where x="(\<lambda>n\<Colon>nat. a $ n \<noteq> b $ n)"])
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   464
  apply (rule ext)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   465
  apply auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   466
  done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   467
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   468
instance ..
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   469
30746
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   470
end
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   471
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   472
lemma fps_nonzero_least_unique:
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   473
  assumes a0: "a \<noteq> 0"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   474
  shows "\<exists>! n. leastP (\<lambda>n. a$n \<noteq> 0) n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   475
proof -
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   476
  from fps_nonzero_nth_minimal [of a] a0
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   477
  obtain n where "a$n \<noteq> 0" "\<forall>m < n. a$m = 0" by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   478
  then have ln: "leastP (\<lambda>n. a$n \<noteq> 0) n"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   479
    by (auto simp add: leastP_def setge_def not_le [symmetric])
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   480
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   481
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   482
    fix m
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   483
    assume "leastP (\<lambda>n. a $ n \<noteq> 0) m"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   484
    then have "m = n" using ln
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   485
      apply (auto simp add: leastP_def setge_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   486
      apply (erule allE[where x=n])
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   487
      apply (erule allE[where x=m])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   488
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   489
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   490
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   491
  ultimately show ?thesis by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   492
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   493
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   494
lemma fps_eq_least_unique:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   495
  assumes ab: "(a::('a::ab_group_add) fps) \<noteq> b"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   496
  shows "\<exists>! n. leastP (\<lambda>n. a$n \<noteq> b$n) n"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   497
  using fps_nonzero_least_unique[of "a - b"] ab
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   498
  by auto
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   499
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   500
instantiation fps :: (comm_ring_1) metric_space
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   501
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   502
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   503
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
   504
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   505
instance
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   506
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   507
  fix S :: "'a fps set"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   508
  show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   509
    by (auto simp add: open_fps_def ball_def subset_eq)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   510
next
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   511
  {
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   512
    fix a b :: "'a fps"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   513
    {
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   514
      assume "a = b"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   515
      then have "\<not> (\<exists>n. a $ n \<noteq> b $ n)" by simp
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   516
      then have "dist a b = 0" by (simp add: dist_fps_def)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   517
    }
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   518
    moreover
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   519
    {
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   520
      assume d: "dist a b = 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   521
      then have "\<forall>n. a$n = b$n"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   522
        by - (rule ccontr, simp add: dist_fps_def)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   523
      then have "a = b" by (simp add: fps_eq_iff)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   524
    }
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   525
    ultimately show "dist a b =0 \<longleftrightarrow> a = b" by blast
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   526
  }
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   527
  note th = this
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   528
  from th have th'[simp]: "\<And>a::'a fps. dist a a = 0" by simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   529
  fix a b c :: "'a fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   530
  {
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   531
    assume "a = b"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   532
    then have "dist a b = 0" unfolding th .
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   533
    then have "dist a b \<le> dist a c + dist b c"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   534
      using dist_fps_ge0 [of a c] dist_fps_ge0 [of b c] by simp
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   535
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   536
  moreover
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   537
  {
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   538
    assume "c = a \<or> c = b"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   539
    then have "dist a b \<le> dist a c + dist b c"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   540
      by (cases "c = a") (simp_all add: th dist_fps_sym)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   541
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   542
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   543
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   544
    assume ab: "a \<noteq> b" and ac: "a \<noteq> c" and bc: "b \<noteq> c"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   545
    let ?P = "\<lambda>a b n. a$n \<noteq> b$n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   546
    from fps_eq_least_unique[OF ab] fps_eq_least_unique[OF ac]
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   547
      fps_eq_least_unique[OF bc]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   548
    obtain nab nac nbc where nab: "leastP (?P a b) nab"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   549
      and nac: "leastP (?P a c) nac"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   550
      and nbc: "leastP (?P b c) nbc" by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   551
    from nab have nab': "\<And>m. m < nab \<Longrightarrow> a$m = b$m" "a$nab \<noteq> b$nab"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   552
      by (auto simp add: leastP_def setge_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   553
    from nac have nac': "\<And>m. m < nac \<Longrightarrow> a$m = c$m" "a$nac \<noteq> c$nac"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   554
      by (auto simp add: leastP_def setge_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   555
    from nbc have nbc': "\<And>m. m < nbc \<Longrightarrow> b$m = c$m" "b$nbc \<noteq> c$nbc"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   556
      by (auto simp add: leastP_def setge_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   557
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   558
    have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   559
      by (simp add: fps_eq_iff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   560
    from ab ac bc nab nac nbc
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   561
    have dab: "dist a b = inverse (2 ^ nab)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   562
      and dac: "dist a c = inverse (2 ^ nac)"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   563
      and dbc: "dist b c = inverse (2 ^ nbc)"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   564
      unfolding th0
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   565
      apply (simp_all add: dist_fps_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   566
      apply (erule the1_equality[OF fps_eq_least_unique[OF ab]])
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   567
      apply (erule the1_equality[OF fps_eq_least_unique[OF ac]])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   568
      apply (erule the1_equality[OF fps_eq_least_unique[OF bc]])
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   569
      done
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   570
    from ab ac bc have nz: "dist a b \<noteq> 0" "dist a c \<noteq> 0" "dist b c \<noteq> 0"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   571
      unfolding th by simp_all
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   572
    from nz have pos: "dist a b > 0" "dist a c > 0" "dist b c > 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   573
      using dist_fps_ge0[of a b] dist_fps_ge0[of a c] dist_fps_ge0[of b c]
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   574
      by auto
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   575
    have th1: "\<And>n. (2::real)^n >0" by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   576
    {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   577
      assume h: "dist a b > dist a c + dist b c"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   578
      then have gt: "dist a b > dist a c" "dist a b > dist b c"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   579
        using pos by auto
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   580
      from gt have gtn: "nab < nbc" "nab < nac"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   581
        unfolding dab dbc dac by (auto simp add: th1)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   582
      from nac'(1)[OF gtn(2)] nbc'(1)[OF gtn(1)]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   583
      have "a $ nab = b $ nab" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   584
      with nab'(2) have False  by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   585
    }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   586
    then have "dist a b \<le> dist a c + dist b c"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   587
      by (auto simp add: not_le[symmetric])
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   588
  }
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   589
  ultimately show "dist a b \<le> dist a c + dist b c" by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   590
qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   591
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   592
end
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   593
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   594
text{* The infinite sums and justification of the notation in textbooks*}
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   595
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   596
lemma reals_power_lt_ex:
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   597
  assumes xp: "x > 0" and y1: "(y::real) > 1"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   598
  shows "\<exists>k>0. (1/y)^k < x"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   599
proof -
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   600
  have yp: "y > 0" using y1 by simp
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   601
  from reals_Archimedean2[of "max 0 (- log y x) + 1"]
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   602
  obtain k::nat where k: "real k > max 0 (- log y x) + 1" by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   603
  from k have kp: "k > 0" by simp
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   604
  from k have "real k > - log y x" by simp
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   605
  then have "ln y * real k > - ln x" unfolding log_def
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   606
    using ln_gt_zero_iff[OF yp] y1
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
   607
    by (simp add: minus_divide_left field_simps del:minus_divide_left[symmetric])
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   608
  then have "ln y * real k + ln x > 0" by simp
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   609
  then have "exp (real k * ln y + ln x) > exp 0"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   610
    by (simp add: mult_ac)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   611
  then have "y ^ k * x > 1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   612
    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
   613
    by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   614
  then have "x > (1 / y)^k" using yp
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
   615
    by (simp add: field_simps nonzero_power_divide)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   616
  then show ?thesis using kp by blast
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   617
qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   618
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   619
lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   620
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   621
lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else (0::'a::comm_ring_1))"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   622
  by (simp add: X_power_iff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   623
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   624
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   625
lemma fps_sum_rep_nth: "(setsum (%i. fps_const(a$i)*X^i) {0..m})$n =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   626
    (if n \<le> m then a$n else (0::'a::comm_ring_1))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   627
  apply (auto simp add: fps_setsum_nth cond_value_iff cong del: if_weak_cong)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   628
  apply (simp add: setsum_delta')
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   629
  done
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   630
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   631
lemma fps_notation: "(%n. setsum (%i. fps_const(a$i) * X^i) {0..n}) ----> a"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   632
  (is "?s ----> a")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   633
proof -
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   634
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   635
    fix r:: real
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   636
    assume rp: "r > 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   637
    have th0: "(2::real) > 1" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   638
    from reals_power_lt_ex[OF rp th0]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   639
    obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   640
    {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   641
      fix n::nat
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   642
      assume nn0: "n \<ge> n0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   643
      then have thnn0: "(1/2)^n <= (1/2 :: real)^n0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   644
        by (auto intro: power_decreasing)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   645
      {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   646
        assume "?s n = a"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   647
        then have "dist (?s n) a < r"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   648
          unfolding dist_eq_0_iff[of "?s n" a, symmetric]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   649
          using rp by (simp del: dist_eq_0_iff)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   650
      }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   651
      moreover
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   652
      {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   653
        assume neq: "?s n \<noteq> a"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   654
        from fps_eq_least_unique[OF neq]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   655
        obtain k where k: "leastP (\<lambda>i. ?s n $ i \<noteq> a$i) k" by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   656
        have th0: "\<And>(a::'a fps) b. a \<noteq> b \<longleftrightarrow> (\<exists>n. a$n \<noteq> b$n)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   657
          by (simp add: fps_eq_iff)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   658
        from neq have dth: "dist (?s n) a = (1/2)^k"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   659
          unfolding th0 dist_fps_def
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   660
          unfolding the1_equality[OF fps_eq_least_unique[OF neq], OF k]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   661
          by (auto simp add: inverse_eq_divide power_divide)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   662
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   663
        from k have kn: "k > n"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   664
          by (simp add: leastP_def setge_def fps_sum_rep_nth split:split_if_asm)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   665
        then have "dist (?s n) a < (1/2)^n" unfolding dth
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   666
          by (auto intro: power_strict_decreasing)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   667
        also have "\<dots> <= (1/2)^n0" using nn0
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   668
          by (auto intro: power_decreasing)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   669
        also have "\<dots> < r" using n0 by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   670
        finally have "dist (?s n) a < r" .
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   671
      }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   672
      ultimately have "dist (?s n) a < r" by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   673
    }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   674
    then have "\<exists>n0. \<forall> n \<ge> n0. dist (?s n) a < r " by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   675
  }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   676
  then show ?thesis unfolding LIMSEQ_def by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   677
qed
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   678
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
   679
subsection{* Inverses of formal power series *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   680
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   681
declare setsum_cong[fundef_cong]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   682
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   683
instantiation fps :: ("{comm_monoid_add, inverse, times, uminus}") inverse
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   684
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   685
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   686
fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   687
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   688
  "natfun_inverse f 0 = inverse (f$0)"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   689
| "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
   690
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   691
definition
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   692
  fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   693
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   694
definition
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   695
  fps_divide_def: "divide = (\<lambda>(f::'a fps) g. f * inverse g)"
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   696
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   697
instance ..
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   698
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   699
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   700
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   701
lemma fps_inverse_zero [simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   702
  "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
   703
  by (simp add: fps_ext fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   704
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   705
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
   706
  apply (auto simp add: expand_fps_eq fps_inverse_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   707
  apply (case_tac n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   708
  apply auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   709
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   710
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   711
lemma inverse_mult_eq_1 [intro]:
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   712
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   713
  shows "inverse f * f = 1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   714
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   715
  have c: "inverse f * f = f * inverse f" by (simp add: mult_commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   716
  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
   717
    by (simp add: fps_inverse_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   718
  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
   719
    by (simp add: fps_mult_nth fps_inverse_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   720
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   721
    fix n :: nat
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   722
    assume np: "n > 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   723
    from np have eq: "{0..n} = {0} \<union> {1 .. n}" by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   724
    have d: "{0} \<inter> {1 .. n} = {}" by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   725
    from f0 np have th0: "- (inverse f $ n) =
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   726
      (setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   727
      by (cases n) (simp_all add: divide_inverse fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   728
    from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   729
    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
   730
      by (simp add: field_simps)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   731
    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
   732
      unfolding fps_mult_nth ifn ..
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   733
    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
   734
      by (simp add: eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   735
    also have "\<dots> = 0" unfolding th1 ifn by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   736
    finally have "(inverse f * f)$n = 0" unfolding c .
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   737
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   738
  with th0 show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   739
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   740
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   741
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
   742
  by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   743
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   744
lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::field) fps) \<longleftrightarrow> f $0 = 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   745
proof -
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   746
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   747
    assume "f$0 = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   748
    then have "inverse f = 0" by (simp add: fps_inverse_def)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   749
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   750
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   751
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   752
    assume h: "inverse f = 0" and c: "f $0 \<noteq> 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   753
    from inverse_mult_eq_1[OF c] h have False by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   754
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   755
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   756
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   757
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   758
lemma fps_inverse_idempotent[intro]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   759
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   760
  shows "inverse (inverse f) = f"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   761
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   762
  from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   763
  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   764
  have "inverse f * f = inverse f * inverse (inverse f)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   765
    by (simp add: mult_ac)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   766
  then show ?thesis using f0 unfolding mult_cancel_left by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   767
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   768
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   769
lemma fps_inverse_unique:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   770
  assumes f0: "f$0 \<noteq> (0::'a::field)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   771
    and fg: "f*g = 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   772
  shows "inverse f = g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   773
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   774
  from inverse_mult_eq_1[OF f0] fg
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   775
  have th0: "inverse f * f = g * f" by (simp add: mult_ac)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   776
  then show ?thesis using f0  unfolding mult_cancel_right
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   777
    by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   778
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   779
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   780
lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   781
    = Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   782
  apply (rule fps_inverse_unique)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   783
  apply simp
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   784
  apply (simp add: fps_eq_iff fps_mult_nth)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   785
proof clarsimp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   786
  fix n :: nat
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   787
  assume n: "n > 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   788
  let ?f = "\<lambda>i. if n = i then (1\<Colon>'a) else if n - i = 1 then - 1 else 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   789
  let ?g = "\<lambda>i. if i = n then 1 else if i=n - 1 then - 1 else 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   790
  let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   791
  have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   792
    by (rule setsum_cong2) auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   793
  have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   794
    using n apply - by (rule setsum_cong2) auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   795
  have eq: "{0 .. n} = {0.. n - 1} \<union> {n}" by auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   796
  from n have d: "{0.. n - 1} \<inter> {n} = {}" by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   797
  have f: "finite {0.. n - 1}" "finite {n}" by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   798
  show "setsum ?f {0..n} = 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   799
    unfolding th1
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   800
    apply (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   801
    unfolding th2
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   802
    apply (simp add: setsum_delta)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   803
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   804
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   805
29912
f4ac160d2e77 fix spelling
huffman
parents: 29911
diff changeset
   806
subsection{* Formal Derivatives, and the MacLaurin theorem around 0*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   807
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   808
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
   809
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   810
lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n+1)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   811
  by (simp add: fps_deriv_def)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   812
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   813
lemma fps_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   814
  "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   815
    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
   816
  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
   817
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   818
lemma fps_deriv_mult[simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   819
  fixes f :: "('a :: comm_ring_1) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   820
  shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   821
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   822
  let ?D = "fps_deriv"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   823
  { fix n::nat
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   824
    let ?Zn = "{0 ..n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   825
    let ?Zn1 = "{0 .. n + 1}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   826
    let ?f = "\<lambda>i. i + 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   827
    have fi: "inj_on ?f {0..n}" by (simp add: inj_on_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   828
    have eq: "{1.. n+1} = ?f ` {0..n}" by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   829
    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
   830
        of_nat (i+1)* f $ (i+1) * g $ (n - i)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   831
    let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   832
        of_nat i* f $ i * g $ ((n + 1) - i)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   833
    {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   834
      fix k
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   835
      assume k: "k \<in> {0..n}"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   836
      have "?h (k + 1) = ?g k" using k by auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   837
    }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   838
    note th0 = this
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   839
    have eq': "{0..n +1}- {1 .. n+1} = {0}" by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   840
    have s0: "setsum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   841
      setsum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   842
      apply (rule setsum_reindex_cong[where f="\<lambda>i. n + 1 - i"])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   843
      apply (simp add: inj_on_def Ball_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   844
      apply presburger
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   845
      apply (rule set_eqI)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   846
      apply (presburger add: image_iff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   847
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   848
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   849
    have s1: "setsum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   850
      setsum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   851
      apply (rule setsum_reindex_cong[where f="\<lambda>i. n + 1 - i"])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   852
      apply (simp add: inj_on_def Ball_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   853
      apply presburger
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   854
      apply (rule set_eqI)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   855
      apply (presburger add: image_iff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   856
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   857
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   858
    have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   859
      by (simp only: mult_commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   860
    also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   861
      by (simp add: fps_mult_nth setsum_addf[symmetric])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   862
    also have "\<dots> = setsum ?h {1..n+1}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   863
      using th0 setsum_reindex_cong[OF fi eq, of "?g" "?h"] by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   864
    also have "\<dots> = setsum ?h {0..n+1}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   865
      apply (rule setsum_mono_zero_left)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   866
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   867
      apply (simp add: subset_eq)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   868
      unfolding eq'
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   869
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   870
      done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   871
    also have "\<dots> = (fps_deriv (f * g)) $ n"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   872
      apply (simp only: fps_deriv_nth fps_mult_nth setsum_addf)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   873
      unfolding s0 s1
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   874
      unfolding setsum_addf[symmetric] setsum_right_distrib
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   875
      apply (rule setsum_cong2)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   876
      apply (auto simp add: of_nat_diff field_simps)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   877
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   878
    finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   879
  }
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   880
  then show ?thesis unfolding fps_eq_iff by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   881
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   882
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   883
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
   884
  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
   885
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   886
lemma fps_deriv_neg[simp]: "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
   887
  by (simp add: fps_eq_iff fps_deriv_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   888
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   889
lemma fps_deriv_add[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) + g) = fps_deriv f + fps_deriv g"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   890
  using fps_deriv_linear[of 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   891
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   892
lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   893
  unfolding diff_minus by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   894
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   895
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
   896
  by (simp add: fps_ext fps_deriv_def fps_const_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   897
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   898
lemma fps_deriv_mult_const_left[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   899
    "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
   900
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   901
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   902
lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   903
  by (simp add: fps_deriv_def fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   904
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   905
lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   906
  by (simp add: fps_deriv_def fps_eq_iff )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   907
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   908
lemma fps_deriv_mult_const_right[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   909
    "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
   910
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   911
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   912
lemma fps_deriv_setsum:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   913
  "fps_deriv (setsum f S) = setsum (\<lambda>i. fps_deriv (f i :: ('a::comm_ring_1) fps)) S"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   914
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   915
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   916
    assume "\<not> finite S"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   917
    then have ?thesis by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   918
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   919
  moreover
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   920
  {
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   921
    assume fS: "finite S"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   922
    have ?thesis by (induct rule: finite_induct [OF fS]) simp_all
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   923
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   924
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   925
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   926
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   927
lemma fps_deriv_eq_0_iff [simp]:
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   928
  "fps_deriv f = 0 \<longleftrightarrow> (f = fps_const (f$0 :: 'a::{idom,semiring_char_0}))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   929
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   930
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   931
    assume "f = fps_const (f$0)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   932
    then have "fps_deriv f = fps_deriv (fps_const (f$0))" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   933
    then have "fps_deriv f = 0" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   934
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   935
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   936
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   937
    assume z: "fps_deriv f = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   938
    then have "\<forall>n. (fps_deriv f)$n = 0" by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   939
    then have "\<forall>n. f$(n+1) = 0" by (simp del: of_nat_Suc of_nat_add One_nat_def)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   940
    then have "f = fps_const (f$0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   941
      apply (clarsimp simp add: fps_eq_iff fps_const_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   942
      apply (erule_tac x="n - 1" in allE)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   943
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   944
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   945
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   946
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   947
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   948
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   949
lemma fps_deriv_eq_iff:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   950
  fixes f:: "('a::{idom,semiring_char_0}) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   951
  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
   952
proof -
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   953
  have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   954
    by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   955
  also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f-g)$0)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   956
    unfolding fps_deriv_eq_0_iff ..
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
   957
  finally show ?thesis by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   958
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   959
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   960
lemma fps_deriv_eq_iff_ex:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   961
  "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>(c::'a::{idom,semiring_char_0}). f = fps_const c + g)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   962
  apply auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   963
  unfolding fps_deriv_eq_iff
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   964
  apply blast
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   965
  done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   966
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   967
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   968
fun fps_nth_deriv :: "nat \<Rightarrow> ('a::semiring_1) fps \<Rightarrow> 'a fps"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   969
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   970
  "fps_nth_deriv 0 f = f"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   971
| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   972
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   973
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
   974
  by (induct n arbitrary: f) auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   975
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   976
lemma fps_nth_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   977
  "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   978
    fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   979
  by (induct n arbitrary: f g) (auto simp add: fps_nth_deriv_commute)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   980
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   981
lemma fps_nth_deriv_neg[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   982
  "fps_nth_deriv n (- (f:: ('a:: comm_ring_1) fps)) = - (fps_nth_deriv n f)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   983
  by (induct n arbitrary: f) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   984
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   985
lemma fps_nth_deriv_add[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   986
  "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
   987
  using fps_nth_deriv_linear[of n 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   988
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   989
lemma fps_nth_deriv_sub[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   990
  "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   991
  unfolding diff_minus fps_nth_deriv_add by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   992
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   993
lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   994
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   995
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   996
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
   997
  by (induct n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   998
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   999
lemma fps_nth_deriv_const[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1000
  "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1001
  by (cases n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1002
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1003
lemma fps_nth_deriv_mult_const_left[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1004
  "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
  1005
  using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1006
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1007
lemma fps_nth_deriv_mult_const_right[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1008
  "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1009
  using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult_commute)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1010
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1011
lemma fps_nth_deriv_setsum:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1012
  "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
  1013
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1014
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1015
  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1016
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1017
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1018
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1019
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1020
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1021
lemma fps_deriv_maclauren_0:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1022
  "(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
  1023
  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
  1024
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1025
subsection {* Powers*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1026
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1027
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
  1028
  by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1029
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1030
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
  1031
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1032
  case 0
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1033
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1034
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1035
  case (Suc n)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1036
  note h = Suc.hyps[OF `a$0 = 1`]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1037
  show ?case unfolding power_Suc fps_mult_nth
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1038
    using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1039
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1040
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1041
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1042
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
  1043
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1044
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1045
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
  1046
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1047
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1048
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
  1049
  by (induct n) (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1050
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1051
lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::{idom}) \<longleftrightarrow> (n \<noteq> 0 \<and> a$0 = 0)"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1052
  apply (rule iffI)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1053
  apply (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1054
  apply (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1055
  apply (rule startsby_zero_power, simp_all)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1056
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1057
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1058
lemma startsby_zero_power_prefix:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1059
  assumes a0: "a $0 = (0::'a::idom)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1060
  shows "\<forall>n < k. a ^ k $ n = 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1061
  using a0
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1062
proof(induct k rule: nat_less_induct)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1063
  fix k
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1064
  assume H: "\<forall>m<k. a $0 =  0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $0 = (0\<Colon>'a)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1065
  let ?ths = "\<forall>m<k. a ^ k $ m = 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1066
  { assume "k = 0" then have ?ths by simp }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1067
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1068
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1069
    fix l
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1070
    assume k: "k = Suc l"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1071
    {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1072
      fix m
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1073
      assume mk: "m < k"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1074
      {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1075
        assume "m = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1076
        then have "a^k $ m = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1077
          using startsby_zero_power[of a k] k a0 by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1078
      }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1079
      moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1080
      {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1081
        assume m0: "m \<noteq> 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1082
        have "a ^k $ m = (a^l * a) $m" by (simp add: k mult_commute)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1083
        also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1084
        also have "\<dots> = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1085
          apply (rule setsum_0')
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1086
          apply auto
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  1087
          apply (case_tac "x = m")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1088
          using a0 apply simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1089
          apply (rule H[rule_format])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1090
          using a0 k mk apply auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1091
          done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1092
        finally have "a^k $ m = 0" .
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1093
      }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1094
      ultimately have "a^k $ m = 0" by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1095
    }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1096
    then have ?ths by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1097
  }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1098
  ultimately show ?ths by (cases k) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1099
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1100
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1101
lemma startsby_zero_setsum_depends:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1102
  assumes a0: "a $0 = (0::'a::idom)" and kn: "n \<ge> k"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1103
  shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1104
  apply (rule setsum_mono_zero_right)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1105
  using kn apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1106
  apply (rule startsby_zero_power_prefix[rule_format, OF a0])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1107
  apply arith
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1108
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1109
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1110
lemma startsby_zero_power_nth_same:
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1111
  assumes a0: "a$0 = (0::'a::{idom})"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1112
  shows "a^n $ n = (a$1) ^ n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1113
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1114
  case 0
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1115
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1116
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1117
  case (Suc n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1118
  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1119
  also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1120
    by (simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1121
  also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1122
    apply (rule setsum_mono_zero_right)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1123
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1124
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1125
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1126
    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1127
    apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1128
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1129
  also have "\<dots> = a^n $ n * a$1" using a0 by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1130
  finally show ?case using Suc.hyps by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1131
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1132
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1133
lemma fps_inverse_power:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1134
  fixes a :: "('a::{field}) fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1135
  shows "inverse (a^n) = inverse a ^ n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1136
proof -
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1137
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1138
    assume a0: "a$0 = 0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1139
    then have eq: "inverse a = 0" by (simp add: fps_inverse_def)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1140
    { assume "n = 0" hence ?thesis by simp }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1141
    moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1142
    {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1143
      assume n: "n > 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1144
      from startsby_zero_power[OF a0 n] eq a0 n have ?thesis
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1145
        by (simp add: fps_inverse_def)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1146
    }
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1147
    ultimately have ?thesis by blast
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1148
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1149
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1150
  {
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1151
    assume a0: "a$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1152
    have ?thesis
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1153
      apply (rule fps_inverse_unique)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1154
      apply (simp add: a0)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1155
      unfolding power_mult_distrib[symmetric]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1156
      apply (rule ssubst[where t = "a * inverse a" and s= 1])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1157
      apply simp_all
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1158
      apply (subst mult_commute)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1159
      apply (rule inverse_mult_eq_1[OF a0])
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1160
      done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1161
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1162
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1163
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1164
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1165
lemma fps_deriv_power:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1166
    "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1167
  apply (induct n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1168
  apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1169
  apply (case_tac n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1170
  apply (auto simp add: field_simps)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1171
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1172
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1173
lemma fps_inverse_deriv:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1174
  fixes a:: "('a :: field) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1175
  assumes a0: "a$0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1176
  shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1177
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1178
  from inverse_mult_eq_1[OF a0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1179
  have "fps_deriv (inverse a * a) = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1180
  hence "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1181
  hence "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0"  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1182
  with inverse_mult_eq_1[OF a0]
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1183
  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
  1184
    unfolding power2_eq_square
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1185
    apply (simp add: field_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1186
    apply (simp add: mult_assoc[symmetric])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1187
    done
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1188
  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
  1189
      0 - fps_deriv a * (inverse a)\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1190
    by simp
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1191
  then show "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1192
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1193
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1194
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1195
lemma fps_inverse_mult:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1196
  fixes a::"('a :: field) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1197
  shows "inverse (a * b) = inverse a * inverse b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1198
proof -
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1199
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1200
    assume a0: "a$0 = 0" hence ab0: "(a*b)$0 = 0" by (simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1201
    from a0 ab0 have th: "inverse a = 0" "inverse (a*b) = 0" by simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1202
    have ?thesis unfolding th by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1203
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1204
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1205
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1206
    assume b0: "b$0 = 0" hence ab0: "(a*b)$0 = 0" by (simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1207
    from b0 ab0 have th: "inverse b = 0" "inverse (a*b) = 0" by simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1208
    have ?thesis unfolding th by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1209
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1210
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1211
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1212
    assume a0: "a$0 \<noteq> 0" and b0: "b$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1213
    from a0 b0 have ab0:"(a*b) $ 0 \<noteq> 0" by (simp  add: fps_mult_nth)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1214
    from inverse_mult_eq_1[OF ab0]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1215
    have "inverse (a*b) * (a*b) * inverse a * inverse b = 1 * inverse a * inverse b" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1216
    then have "inverse (a*b) * (inverse a * a) * (inverse b * b) = inverse a * inverse b"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1217
      by (simp add: field_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1218
    then have ?thesis using inverse_mult_eq_1[OF a0] inverse_mult_eq_1[OF b0] by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1219
  }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1220
  ultimately show ?thesis by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1221
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1222
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1223
lemma fps_inverse_deriv':
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1224
  fixes a:: "('a :: field) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1225
  assumes a0: "a$0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1226
  shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1227
  using fps_inverse_deriv[OF a0]
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1228
  unfolding power2_eq_square fps_divide_def fps_inverse_mult
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1229
  by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1230
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1231
lemma inverse_mult_eq_1':
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1232
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1233
  shows "f * inverse f= 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1234
  by (metis mult_commute inverse_mult_eq_1 f0)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1235
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1236
lemma fps_divide_deriv:
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1237
  fixes a:: "('a :: field) fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1238
  assumes a0: "b$0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1239
  shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1240
  using fps_inverse_deriv[OF a0]
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1241
  by (simp add: fps_divide_def field_simps
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1242
    power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1243
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1244
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1245
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
  1246
  by (simp add: fps_inverse_gp fps_eq_iff X_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1247
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1248
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
  1249
  by (cases n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1250
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1251
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1252
lemma fps_inverse_X_plus1:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1253
  "inverse (1 + X) = Abs_fps (\<lambda>n. (- (1::'a::{field})) ^ n)" (is "_ = ?r")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1254
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1255
  have eq: "(1 + X) * ?r = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1256
    unfolding minus_one_power_iff
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1257
    by (auto simp add: field_simps fps_eq_iff)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1258
  show ?thesis by (auto simp add: eq intro: fps_inverse_unique simp del: minus_one)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1259
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1260
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1261
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1262
subsection{* Integration *}
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1263
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1264
definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1265
  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
  1266
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1267
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
  1268
  unfolding fps_integral_def fps_deriv_def
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1269
  by (simp add: fps_eq_iff del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1270
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1271
lemma fps_integral_linear:
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1272
  "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
  1273
    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
  1274
  (is "?l = ?r")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1275
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1276
  have "fps_deriv ?l = fps_deriv ?r" by (simp add: fps_deriv_fps_integral)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1277
  moreover have "?l$0 = ?r$0" by (simp add: fps_integral_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1278
  ultimately show ?thesis
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1279
    unfolding fps_deriv_eq_iff by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1280
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1281
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1282
subsection {* Composition of FPSs *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1283
definition fps_compose :: "('a::semiring_1) fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55) where
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1284
  fps_compose_def: "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1285
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1286
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
  1287
  by (simp add: fps_compose_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1288
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1289
lemma fps_compose_X[simp]: "a oo X = (a :: ('a :: comm_ring_1) fps)"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  1290
  by (simp add: fps_ext fps_compose_def mult_delta_right setsum_delta')
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1291
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1292
lemma fps_const_compose[simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1293
  "fps_const (a::'a::{comm_ring_1}) oo b = fps_const (a)"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  1294
  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
  1295
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1296
lemma numeral_compose[simp]: "(numeral k::('a::{comm_ring_1}) fps) oo b = numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1297
  unfolding numeral_fps_const by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1298
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1299
lemma neg_numeral_compose[simp]: "(neg_numeral k::('a::{comm_ring_1}) fps) oo b = neg_numeral k"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1300
  unfolding neg_numeral_fps_const by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  1301
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1302
lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: ('a :: comm_ring_1) fps)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1303
  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
  1304
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1305
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1306
subsection {* Rules from Herbert Wilf's Generatingfunctionology*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1307
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1308
subsubsection {* Rule 1 *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1309
  (* {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
  1310
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1311
lemma fps_power_mult_eq_shift:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1312
  "X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1313
    Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: comm_ring_1) * X^i) {0 .. k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1314
  (is "?lhs = ?rhs")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1315
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1316
  { fix n:: nat
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1317
    have "?lhs $ n = (if n < Suc k then 0 else a n)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1318
      unfolding X_power_mult_nth by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1319
    also have "\<dots> = ?rhs $ n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1320
    proof (induct k)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1321
      case 0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1322
      thus ?case by (simp add: fps_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1323
    next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1324
      case (Suc k)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1325
      note th = Suc.hyps[symmetric]
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1326
      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1327
        (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1328
        by (simp add: field_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1329
      also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1330
        using th unfolding fps_sub_nth by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1331
      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
  1332
        unfolding X_power_mult_right_nth
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1333
        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
  1334
        apply (rule cong[of a a, OF refl])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1335
        apply arith
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1336
        done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1337
      finally show ?case by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1338
    qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1339
    finally have "?lhs $ n = ?rhs $ n" .
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1340
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1341
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1342
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1343
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1344
subsubsection{* Rule 2*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1345
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1346
  (* 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
  1347
  (* If f reprents {a_n} and P is a polynomial, then
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1348
        P(xD) f represents {P(n) a_n}*)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1349
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1350
definition "XD = op * X o fps_deriv"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1351
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1352
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
  1353
  by (simp add: XD_def field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1354
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1355
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
  1356
  by (simp add: XD_def field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1357
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1358
lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1359
    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
  1360
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1361
30952
7ab2716dd93b power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents: 30837
diff changeset
  1362
lemma XDN_linear:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1363
  "(XD ^^ n) (fps_const c * a + fps_const d * b) =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1364
    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
  1365
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1366
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1367
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
  1368
  by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1369
30994
chaieb
parents: 30971 30992
diff changeset
  1370
30952
7ab2716dd93b power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents: 30837
diff changeset
  1371
lemma fps_mult_XD_shift:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1372
  "(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
  1373
  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
  1374
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1375
subsubsection{* Rule 3 is trivial and is given by @{text fps_times_def}*}
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1376
subsubsection{* Rule 5 --- summation and "division" by (1 - X)*}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1377
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1378
lemma fps_divide_X_minus1_setsum_lemma:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1379
  "a = ((1::('a::comm_ring_1) fps) - X) * Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1380
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1381
  let ?X = "X::('a::comm_ring_1) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1382
  let ?sa = "Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1383
  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
  1384
    by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1385
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1386
    fix n:: nat
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1387
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1388
      assume "n=0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1389
      hence "a$n = ((1 - ?X) * ?sa) $ n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1390
        by (simp add: fps_mult_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1391
    }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1392
    moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1393
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1394
      assume n0: "n \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1395
      then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1}\<union>{2..n} = {1..n}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1396
        "{0..n - 1}\<union>{n} = {0..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
  1397
        by (auto simp: set_eq_iff)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1398
      have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1399
        "{0..n - 1}\<inter>{n} ={}" using n0 by simp_all
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1400
      have f: "finite {0}" "finite {1}" "finite {2 .. n}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1401
        "finite {0 .. n - 1}" "finite {n}" by simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1402
      have "((1 - ?X) * ?sa) $ n = setsum (\<lambda>i. (1 - ?X)$ i * ?sa $ (n - i)) {0 .. n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1403
        by (simp add: fps_mult_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1404
      also have "\<dots> = a$n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1405
        unfolding th0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1406
        unfolding setsum_Un_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1407
        unfolding setsum_Un_disjoint[OF f(2) f(3) d(2)]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1408
        apply (simp)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1409
        unfolding setsum_Un_disjoint[OF f(4,5) d(3), unfolded u(3)]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1410
        apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1411
        done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1412
      finally have "a$n = ((1 - ?X) * ?sa) $ n" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1413
    }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1414
    ultimately have "a$n = ((1 - ?X) * ?sa) $ n" by blast
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1415
  }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1416
  then show ?thesis unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1417
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1418
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1419
lemma fps_divide_X_minus1_setsum:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1420
  "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
  1421
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1422
  let ?X = "1 - (X::('a::field) fps)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1423
  have th0: "?X $ 0 \<noteq> 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1424
  have "a /?X = ?X *  Abs_fps (\<lambda>n\<Colon>nat. setsum (op $ a) {0..n}) * inverse ?X"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1425
    using fps_divide_X_minus1_setsum_lemma[of a, symmetric] th0
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1426
    by (simp add: fps_divide_def mult_assoc)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1427
  also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n\<Colon>nat. setsum (op $ a) {0..n}) "
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1428
    by (simp add: mult_ac)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1429
  finally show ?thesis by (simp add: inverse_mult_eq_1[OF th0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1430
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1431
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1432
subsubsection{* Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1433
  finite product of FPS, also the relvant instance of powers of a FPS*}
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1434
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1435
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
  1436
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1437
lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1438
  apply (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1439
  apply (case_tac x)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1440
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1441
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1442
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1443
lemma append_natpermute_less_eq:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1444
  assumes h: "xs@ys \<in> natpermute n k"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1445
  shows "listsum xs \<le> n" and "listsum ys \<le> n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1446
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1447
  from h have "listsum (xs @ ys) = n" by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1448
  hence *: "listsum xs + listsum ys = n" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1449
  from * show "listsum xs \<le> n" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1450
  from * show "listsum ys \<le> n" by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1451
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1452
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1453
lemma natpermute_split:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1454
  assumes mn: "h \<le> k"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1455
  shows "natpermute n k =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1456
    (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1457
  (is "?L = ?R" is "?L = (\<Union>m \<in>{0..n}. ?S m)")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1458
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1459
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1460
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1461
    assume l: "l \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1462
    from l obtain m xs ys where h: "m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1463
      and xs: "xs \<in> natpermute m h"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1464
      and ys: "ys \<in> natpermute (n - m) (k - h)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1465
      and leq: "l = xs@ys" by blast
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1466
    from xs have xs': "listsum xs = m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1467
      by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1468
    from ys have ys': "listsum ys = n - m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1469
      by (simp add: natpermute_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1470
    have "l \<in> ?L" using leq xs ys h
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1471
      apply (clarsimp simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1472
      unfolding xs' ys'
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1473
      using mn xs ys
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1474
      unfolding natpermute_def
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1475
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1476
      done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1477
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1478
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1479
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1480
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1481
    assume l: "l \<in> natpermute n k"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1482
    let ?xs = "take h l"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1483
    let ?ys = "drop h l"
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1484
    let ?m = "listsum ?xs"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1485
    from l have ls: "listsum (?xs @ ?ys) = n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1486
      by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1487
    have xs: "?xs \<in> natpermute ?m h" using l mn
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1488
      by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1489
    have l_take_drop: "listsum l = listsum (take h l @ drop h l)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1490
      by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1491
    then have ys: "?ys \<in> natpermute (n - ?m) (k - h)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1492
      using l mn ls by (auto simp add: natpermute_def simp del: append_take_drop_id)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1493
    from ls have m: "?m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1494
      by (simp add: l_take_drop del: append_take_drop_id)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1495
    from xs ys ls have "l \<in> ?R"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1496
      apply auto
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1497
      apply (rule bexI [where x = "?m"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1498
      apply (rule exI [where x = "?xs"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1499
      apply (rule exI [where x = "?ys"])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1500
      using ls l
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1501
      apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1502
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1503
      done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1504
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1505
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1506
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1507
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1508
lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1509
  by (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1510
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1511
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
  1512
  apply (auto simp add: set_replicate_conv_if natpermute_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1513
  apply (rule nth_equalityI)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1514
  apply simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1515
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1516
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1517
lemma natpermute_finite: "finite (natpermute n k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1518
proof (induct k arbitrary: n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1519
  case 0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1520
  then show ?case
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1521
    apply (subst natpermute_split[of 0 0, simplified])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1522
    apply (simp add: natpermute_0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1523
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1524
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1525
  case (Suc k)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1526
  then show ?case unfolding natpermute_split [of k "Suc k", simplified]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1527
    apply -
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1528
    apply (rule finite_UN_I)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1529
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1530
    unfolding One_nat_def[symmetric] natlist_trivial_1
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1531
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1532
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1533
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1534
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1535
lemma natpermute_contain_maximal:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1536
  "{xs \<in> natpermute n (k+1). n \<in> set xs} = UNION {0 .. k} (\<lambda>i. {(replicate (k+1) 0) [i:=n]})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1537
  (is "?A = ?B")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1538
proof -
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1539
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1540
    fix xs
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1541
    assume H: "xs \<in> natpermute n (k+1)" and n: "n \<in> set xs"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1542
    from n obtain i where i: "i \<in> {0.. k}" "xs!i = n" using H
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1543
      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1544
    have eqs: "({0..k} - {i}) \<union> {i} = {0..k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1545
      using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1546
    have f: "finite({0..k} - {i})" "finite {i}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1547
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1548
    have d: "({0..k} - {i}) \<inter> {i} = {}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1549
      using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1550
    from H have "n = setsum (nth xs) {0..k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1551
      apply (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1552
      apply (auto simp add: atLeastLessThanSuc_atLeastAtMost listsum_setsum_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1553
      done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1554
    also have "\<dots> = n + setsum (nth xs) ({0..k} - {i})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1555
      unfolding setsum_Un_disjoint[OF f d, unfolded eqs] using i by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1556
    finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1557
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1558
    from H have xsl: "length xs = k+1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1559
      by (simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1560
    from i have i': "i < length (replicate (k+1) 0)"   "i < k+1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1561
      unfolding length_replicate by presburger+
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1562
    have "xs = replicate (k+1) 0 [i := n]"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1563
      apply (rule nth_equalityI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1564
      unfolding xsl length_list_update length_replicate
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1565
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1566
      apply clarify
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1567
      unfolding nth_list_update[OF i'(1)]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1568
      using i zxs
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1569
      apply (case_tac "ia = i")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1570
      apply (auto simp del: replicate.simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1571
      done
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1572
    then have "xs \<in> ?B" using i by blast
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1573
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1574
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1575
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1576
    fix i
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1577
    assume i: "i \<in> {0..k}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1578
    let ?xs = "replicate (k+1) 0 [i:=n]"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1579
    have nxs: "n \<in> set ?xs"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1580
      apply (rule set_update_memI)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1581
      using i apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1582
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1583
    have xsl: "length ?xs = k+1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1584
      by (simp only: length_replicate length_list_update)
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1585
    have "listsum ?xs = setsum (nth ?xs) {0..<k+1}"
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1586
      unfolding listsum_setsum_nth xsl ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1587
    also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1588
      by (rule setsum_cong2) (simp del: replicate.simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1589
    also have "\<dots> = n" using i by (simp add: setsum_delta)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1590
    finally have "?xs \<in> natpermute n (k+1)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1591
      using xsl unfolding natpermute_def mem_Collect_eq by blast
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1592
    then have "?xs \<in> ?A"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1593
      using nxs  by blast
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1594
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1595
  ultimately show ?thesis by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1596
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1597
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1598
    (* The general form *)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1599
lemma fps_setprod_nth:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1600
  fixes m :: nat
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1601
    and a :: "nat \<Rightarrow> ('a::comm_ring_1) fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1602
  shows "(setprod a {0 .. m})$n = setsum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1603
  (is "?P m n")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1604
proof (induct m arbitrary: n rule: nat_less_induct)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1605
  fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1606
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1607
    assume m0: "m = 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1608
    hence "?P m n" apply simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1609
      unfolding natlist_trivial_1[where n = n, unfolded One_nat_def] by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1610
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1611
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1612
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1613
    fix k assume k: "m = Suc k"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1614
    have km: "k < m" using k by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1615
    have u0: "{0 .. k} \<union> {m} = {0..m}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1616
      using k apply (simp add: set_eq_iff)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1617
      apply presburger
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1618
      done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1619
    have f0: "finite {0 .. k}" "finite {m}" by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1620
    have d0: "{0 .. k} \<inter> {m} = {}" using k by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1621
    have "(setprod a {0 .. m}) $ n = (setprod a {0 .. k} * a m) $ n"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1622
      unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1623
    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
  1624
      unfolding fps_mult_nth H[rule_format, OF km] ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1625
    also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1626
      apply (simp add: k)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1627
      unfolding natpermute_split[of m "m + 1", simplified, of n,
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1628
        unfolded natlist_trivial_1[unfolded One_nat_def] k]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1629
      apply (subst setsum_UN_disjoint)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1630
      apply simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1631
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1632
      unfolding image_Collect[symmetric]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1633
      apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1634
      apply (rule finite_imageI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1635
      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
  1636
      apply (clarsimp simp add: set_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1637
      apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1638
      apply (rule setsum_cong2)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1639
      unfolding setsum_left_distrib
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1640
      apply (rule sym)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1641
      apply (rule_tac f="\<lambda>xs. xs @[n - x]" in  setsum_reindex_cong)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1642
      apply (simp add: inj_on_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1643
      apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1644
      unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0, unfolded k]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1645
      apply (clarsimp simp add: natpermute_def nth_append)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1646
      done
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1647
    finally have "?P m n" .
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1648
  }
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1649
  ultimately show "?P m n " by (cases m) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1650
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1651
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1652
text{* The special form for powers *}
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1653
lemma fps_power_nth_Suc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1654
  fixes m :: nat
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1655
    and a :: "('a::comm_ring_1) fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1656
  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
  1657
proof -
46757
ad878aff9c15 removing finiteness goals
bulwahn
parents: 46131
diff changeset
  1658
  have th0: "a^Suc m = setprod (\<lambda>i. a) {0..m}" by (simp add: setprod_constant)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1659
  show ?thesis unfolding th0 fps_setprod_nth ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1660
qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1661
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1662
lemma fps_power_nth:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1663
  fixes m :: nat and a :: "('a::comm_ring_1) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1664
  shows "(a ^m)$n = (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
  1665
  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
  1666
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1667
lemma fps_nth_power_0:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1668
  fixes m :: nat and a :: "('a::{comm_ring_1}) fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1669
  shows "(a ^m)$0 = (a$0) ^ m"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1670
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1671
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1672
    assume "m = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1673
    hence ?thesis by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1674
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1675
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1676
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1677
    fix n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1678
    assume m: "m = Suc n"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1679
    have c: "m = card {0..n}" using m by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1680
    have "(a ^m)$0 = setprod (\<lambda>i. a$0) {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1681
      by (simp add: m fps_power_nth del: replicate.simps power_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1682
    also have "\<dots> = (a$0) ^ m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1683
     unfolding c by (rule setprod_constant) simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1684
   finally have ?thesis .
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1685
  }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1686
  ultimately show ?thesis by (cases m) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1687
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1688
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1689
lemma fps_compose_inj_right:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  1690
  assumes a0: "a$0 = (0::'a::{idom})"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1691
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1692
  shows "(b oo a = c oo a) \<longleftrightarrow> b = c" (is "?lhs \<longleftrightarrow>?rhs")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1693
proof-
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1694
  { assume ?rhs then have "?lhs" by simp }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1695
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1696
  { assume h: ?lhs
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1697
    { fix n have "b$n = c$n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1698
      proof (induct n rule: nat_less_induct)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1699
        fix n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1700
        assume H: "\<forall>m<n. b$m = c$m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1701
        {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1702
          assume n0: "n=0"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1703
          from h have "(b oo a)$n = (c oo a)$n" by simp
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1704
          hence "b$n = c$n" using n0 by (simp add: fps_compose_nth)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1705
        }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1706
        moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1707
        {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1708
          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
  1709
          have f: "finite {0 .. n1}" "finite {n}" by simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1710
          have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using n1 by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1711
          have d: "{0 .. n1} \<inter> {n} = {}" using n1 by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1712
          have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1713
            apply (rule setsum_cong2)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1714
            using H n1 apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1715
            done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1716
          have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1717
            unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq] seq
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1718
            using startsby_zero_power_nth_same[OF a0]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1719
            by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1720
          have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1721
            unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1722
            using startsby_zero_power_nth_same[OF a0]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1723
            by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1724
          from h[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1725
          have "b$n = c$n" by auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1726
        }
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1727
        ultimately show "b$n = c$n" by (cases n) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1728
      qed}
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1729
    then have ?rhs by (simp add: fps_eq_iff)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1730
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1731
  ultimately show ?thesis by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1732
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1733
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1734
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  1735
subsection {* Radicals *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1736
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1737
declare setprod_cong [fundef_cong]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1738
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1739
function radical :: "(nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> ('a::{field}) fps \<Rightarrow> nat \<Rightarrow> 'a"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1740
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1741
  "radical r 0 a 0 = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1742
| "radical r 0 a (Suc n) = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1743
| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1744
| "radical r (Suc k) a (Suc n) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1745
    (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
  1746
      {xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1747
    (of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1748
  by pat_completeness auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1749
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1750
termination radical
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1751
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1752
  let ?R = "measure (\<lambda>(r, k, a, n). n)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1753
  {
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1754
    show "wf ?R" by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1755
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1756
    fix r k a n xs i
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1757
    assume xs: "xs \<in> {xs \<in> natpermute (Suc n) (Suc k). Suc n \<notin> set xs}" and i: "i \<in> {0..k}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1758
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1759
      assume c: "Suc n \<le> xs ! i"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1760
      from xs i have "xs !i \<noteq> Suc n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1761
        by (auto simp add: in_set_conv_nth natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1762
      with c have c': "Suc n < xs!i" by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1763
      have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1764
        by simp_all
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1765
      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
  1766
        by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1767
      have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1768
        using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1769
      from xs have "Suc n = listsum xs"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1770
        by (simp add: natpermute_def)
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1771
      also have "\<dots> = setsum (nth xs) {0..<Suc k}" using xs
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  1772
        by (simp add: natpermute_def listsum_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1773
      also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1774
        unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1775
        unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1776
        by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1777
      finally have False using c' by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1778
    }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1779
    then show "((r, Suc k, a, xs!i), r, Suc k, a, Suc n) \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1780
      apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1781
      apply (metis not_less)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1782
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1783
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1784
    fix r k a n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1785
    show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \<in> ?R" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1786
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1787
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1788
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1789
definition "fps_radical r n a = Abs_fps (radical r n a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1790
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1791
lemma fps_radical0[simp]: "fps_radical r 0 a = 1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1792
  apply (auto simp add: fps_eq_iff fps_radical_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1793
  apply (case_tac n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1794
  apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1795
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1796
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1797
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
  1798
  by (cases n) (simp_all add: fps_radical_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1799
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1800
lemma fps_radical_power_nth[simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1801
  assumes r: "(r k (a$0)) ^ k = a$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1802
  shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1803
proof-
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1804
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1805
    assume "k = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1806
    hence ?thesis by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1807
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1808
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1809
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1810
    fix h
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1811
    assume h: "k = Suc h"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1812
    have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1813
      unfolding fps_power_nth h by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1814
    also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1815
      apply (rule setprod_cong)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1816
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1817
      using h
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1818
      apply (subgoal_tac "replicate k (0::nat) ! x = 0")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1819
      apply (auto intro: nth_replicate simp del: replicate.simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1820
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1821
    also have "\<dots> = a$0" using r by (simp add: h setprod_constant)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1822
    finally have ?thesis using h by simp}
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1823
  ultimately show ?thesis by (cases k) auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1824
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1825
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1826
lemma natpermute_max_card:
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1827
  assumes n0: "n\<noteq>0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1828
  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
  1829
  unfolding natpermute_contain_maximal
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1830
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1831
  let ?A= "\<lambda>i. {replicate (k + 1) 0[i := n]}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1832
  let ?K = "{0 ..k}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1833
  have fK: "finite ?K" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1834
  have fAK: "\<forall>i\<in>?K. finite (?A i)" by auto
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1835
  have d: "\<forall>i\<in> ?K. \<forall>j\<in> ?K. i \<noteq> j \<longrightarrow>
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1836
    {replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1837
  proof clarify
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1838
    fix i j
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1839
    assume i: "i \<in> ?K" and j: "j\<in> ?K" and ij: "i\<noteq>j"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1840
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1841
      assume eq: "replicate (k+1) 0 [i:=n] = replicate (k+1) 0 [j:= n]"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1842
      have "(replicate (k+1) 0 [i:=n] ! i) = n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1843
        using i by (simp del: replicate.simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1844
      moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1845
      have "(replicate (k+1) 0 [j:=n] ! i) = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1846
        using i ij by (simp del: replicate.simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1847
      ultimately have False
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1848
        using eq n0 by (simp del: replicate.simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1849
    }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1850
    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
  1851
      by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1852
  qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1853
  from card_UN_disjoint[OF fK fAK d] show "card (\<Union>i\<in>{0..k}. {replicate (k + 1) 0[i := n]}) = k+1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1854
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1855
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1856
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1857
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1858
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1859
  assumes a0: "a$0 \<noteq> 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1860
  shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \<longleftrightarrow> (fps_radical r (Suc k) a) ^ (Suc k) = a"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1861
proof-
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1862
  let ?r = "fps_radical r (Suc k) a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1863
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1864
    assume r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1865
    from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1866
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1867
      fix z
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1868
      have "?r ^ Suc k $ z = a$z"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1869
      proof (induct z rule: nat_less_induct)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1870
        fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1871
        assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1872
        {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1873
          assume "n = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1874
          hence "?r ^ Suc k $ n = a $n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1875
            using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1876
        }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1877
        moreover
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1878
        {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1879
          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
  1880
          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
  1881
          let ?Pnk = "natpermute n (k + 1)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1882
          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
  1883
          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
  1884
          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
  1885
          have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1886
          have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1887
            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
  1888
            by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1889
          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
  1890
          have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1891
          proof (rule setsum_cong2)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1892
            fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1893
            let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1894
              fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1895
            from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1896
              unfolding natpermute_contain_maximal by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1897
            have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1898
                (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1899
              apply (rule setprod_cong, simp)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1900
              using i r0 apply (simp del: replicate.simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1901
              done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1902
            also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1903
              using i r0 by (simp add: setprod_gen_delta)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1904
            finally show ?ths .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1905
          qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1906
          then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1907
            by (simp add: natpermute_max_card[OF nz, simplified])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1908
          also have "\<dots> = a$n - setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1909
            unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1910
          finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1911
          have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1912
            unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1913
          also have "\<dots> = a$n" unfolding fn by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1914
          finally have "?r ^ Suc k $ n = a $n" .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1915
        }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1916
        ultimately  show "?r ^ Suc k $ n = a $n" by (cases n) auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1917
      qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1918
    }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1919
    then have ?thesis using r0 by (simp add: fps_eq_iff)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1920
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1921
  moreover
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1922
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1923
    assume h: "(fps_radical r (Suc k) a) ^ (Suc k) = a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1924
    hence "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1925
    then have "(r (Suc k) (a$0)) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1926
      unfolding fps_power_nth_Suc
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1927
      by (simp add: setprod_constant del: replicate.simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1928
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1929
  ultimately show ?thesis by blast
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1930
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1931
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1932
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1933
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1934
  fixes a:: "'a::field_char_0 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1935
  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
  1936
  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1937
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1938
  let ?r = "fps_radical r (Suc k) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1939
  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
  1940
  {fix z have "?r ^ Suc k $ z = a$z"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1941
    proof(induct z rule: nat_less_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1942
      fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1943
      {assume "n = 0" hence "?r ^ Suc k $ n = a $n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1944
          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
  1945
      moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1946
      {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
  1947
        have fK: "finite {0..k}" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1948
        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
  1949
        let ?Pnk = "natpermute n (k + 1)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1950
        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
  1951
        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
  1952
        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
  1953
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1954
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1955
          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
  1956
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1957
        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
  1958
        have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1959
        proof(rule setsum_cong2)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1960
          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
  1961
          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
  1962
          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
  1963
            unfolding natpermute_contain_maximal by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1964
          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))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1965
            apply (rule setprod_cong, simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1966
            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
  1967
          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
  1968
            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
  1969
          finally show ?ths .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1970
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1971
        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
  1972
          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
  1973
        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
  1974
          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
  1975
        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
  1976
        have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1977
          unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1978
        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
  1979
        finally have "?r ^ Suc k $ n = a $n" .}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1980
      ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1981
  qed }
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1982
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1983
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1984
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  1985
*)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1986
lemma eq_divide_imp':
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1987
  assumes c0: "(c::'a::field) ~= 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1988
    and eq: "a * c = b"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1989
  shows "a = b / c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1990
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1991
  from eq have "a * c * inverse c = b * inverse c"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1992
    by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1993
  hence "a * (inverse c * c) = b/c"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1994
    by (simp only: field_simps divide_inverse)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1995
  then show "a = b/c"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1996
    unfolding  field_inverse[OF c0] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1997
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1998
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1999
lemma radical_unique:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2000
  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2001
    and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2002
    and b0: "b$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2003
  shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2004
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2005
  let ?r = "fps_radical r (Suc k) b"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2006
  have r00: "r (Suc k) (b$0) \<noteq> 0" using b0 r0 by auto
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2007
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2008
    assume H: "a = ?r"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2009
    from H have "a^Suc k = b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2010
      using power_radical[OF b0, of r k, unfolded r0] by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2011
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2012
  moreover
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2013
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2014
    assume H: "a^Suc k = b"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2015
    have ceq: "card {0..k} = Suc k" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2016
    from a0 have a0r0: "a$0 = ?r$0" by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2017
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2018
      fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2019
      have "a $ n = ?r $ n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2020
      proof (induct n rule: nat_less_induct)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2021
        fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2022
        assume h: "\<forall>m<n. a$m = ?r $m"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2023
        {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2024
          assume "n = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2025
          hence "a$n = ?r $n" using a0 by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2026
        }
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2027
        moreover
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2028
        {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2029
          fix n1
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2030
          assume n1: "n = Suc n1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2031
          have fK: "finite {0..k}" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2032
        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
  2033
        let ?Pnk = "natpermute n (Suc k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2034
        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
  2035
        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
  2036
        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
  2037
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2038
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2039
          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
  2040
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2041
        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
  2042
        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
  2043
        have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2044
        proof (rule setsum_cong2)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2045
          fix v
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2046
          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
  2047
          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
  2048
          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
  2049
            unfolding Suc_eq_plus1 natpermute_contain_maximal
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2050
            by (auto simp del: replicate.simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2051
          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))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2052
            apply (rule setprod_cong, simp)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2053
            using i a0 apply (simp del: replicate.simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2054
            done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2055
          also have "\<dots> = a $ n * (?r $ 0)^k"
46757
ad878aff9c15 removing finiteness goals
bulwahn
parents: 46131
diff changeset
  2056
            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
  2057
          finally show ?ths .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2058
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2059
        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
  2060
          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
  2061
        have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2062
        proof (rule setsum_cong2, rule setprod_cong, simp)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2063
          fix xs i
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2064
          assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2065
          {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2066
            assume c: "n \<le> xs ! i"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2067
            from xs i have "xs !i \<noteq> n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2068
              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
  2069
            with c have c': "n < xs!i" by arith
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2070
            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2071
              by simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2072
            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
  2073
              by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2074
            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2075
              using i by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2076
            from xs have "n = listsum xs"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2077
              by (simp add: natpermute_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2078
            also have "\<dots> = setsum (nth xs) {0..<Suc k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2079
              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
  2080
            also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2081
              unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2082
              unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2083
              by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2084
            finally have False using c' by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2085
          }
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2086
          then have thn: "xs!i < n" by presburger
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2087
          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
  2088
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2089
        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
  2090
          by (simp add: field_simps del: of_nat_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2091
        from H have "b$n = a^Suc k $ n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2092
          by (simp add: fps_eq_iff)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2093
        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
  2094
          unfolding fps_power_nth_Suc
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2095
          using setsum_Un_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2096
            unfolded eq, of ?g] by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2097
        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2098
          unfolding th0 th1 ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2099
        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2100
          by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2101
        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
  2102
          apply -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2103
          apply (rule eq_divide_imp')
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2104
          using r00
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2105
          apply (simp del: of_nat_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2106
          apply (simp add: mult_ac)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2107
          done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2108
        then have "a$n = ?r $n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2109
          apply (simp del: of_nat_Suc)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2110
          unfolding fps_radical_def n1
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2111
          apply (simp add: field_simps n1 th00 del: of_nat_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2112
          done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2113
        }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2114
        ultimately show "a$n = ?r $ n" by (cases n) auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2115
      qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2116
    }
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2117
    then have "a = ?r" by (simp add: fps_eq_iff)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2118
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2119
  ultimately show ?thesis by blast
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
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2123
lemma radical_power:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2124
  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2125
    and a0: "(a$0 ::'a::field_char_0) \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2126
  shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2127
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2128
  let ?ak = "a^ Suc k"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2129
  have ak0: "?ak $ 0 = (a$0) ^ Suc k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2130
    by (simp add: fps_nth_power_0 del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2131
  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
  2132
    using ak0 by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2133
  from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2134
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2135
  from ak0 a0 have ak00: "?ak $ 0 \<noteq>0 "
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2136
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2137
  from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2138
    by metis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2139
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2140
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2141
lemma fps_deriv_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2142
  fixes a:: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2143
  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2144
    and a0: "a$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2145
  shows "fps_deriv (fps_radical r (Suc k) a) = 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
  2146
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2147
  let ?r = "fps_radical r (Suc k) a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2148
  let ?w = "(fps_const (of_nat (Suc k)) * ?r ^ k)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2149
  from a0 r0 have r0': "r (Suc k) (a$0) \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2150
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2151
  from r0' have w0: "?w $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2152
    by (simp del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2153
  note th0 = inverse_mult_eq_1[OF w0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2154
  let ?iw = "inverse ?w"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2155
  from iffD1[OF power_radical[of a r], OF a0 r0]
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2156
  have "fps_deriv (?r ^ Suc k) = fps_deriv a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2157
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2158
  hence "fps_deriv ?r * ?w = fps_deriv a"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 29915
diff changeset
  2159
    by (simp add: fps_deriv_power mult_ac del: power_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2160
  hence "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2161
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2162
  hence "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2163
    by (simp add: fps_divide_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2164
  then show ?thesis unfolding th0 by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2165
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2166
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2167
lemma radical_mult_distrib:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2168
  fixes a:: "'a::field_char_0 fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2169
  assumes k: "k > 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2170
    and ra0: "r k (a $ 0) ^ k = a $ 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2171
    and rb0: "r k (b $ 0) ^ k = b $ 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2172
    and a0: "a$0 \<noteq> 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2173
    and b0: "b$0 \<noteq> 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2174
  shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \<longleftrightarrow>
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2175
    fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2176
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2177
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2178
    assume  r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2179
    from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2180
      by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2181
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2182
      assume "k = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2183
      hence ?thesis using r0' by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2184
    }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2185
    moreover
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2186
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2187
      fix h assume k: "k = Suc h"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2188
      let ?ra = "fps_radical r (Suc h) a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2189
      let ?rb = "fps_radical r (Suc h) b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2190
      have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2191
        using r0' k by (simp add: fps_mult_nth)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2192
      have ab0: "(a*b) $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2193
        using a0 b0 by (simp add: fps_mult_nth)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2194
      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]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2195
        iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded k]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded k]] k r0'
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2196
      have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2197
    }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2198
    ultimately have ?thesis by (cases k) auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2199
  }
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2200
  moreover
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2201
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2202
    assume h: "fps_radical r k (a*b) = fps_radical r k a * fps_radical r k b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2203
    hence "(fps_radical r k (a*b))$0 = (fps_radical r k a * fps_radical r k b)$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2204
      by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2205
    then have "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2206
      using k by (simp add: fps_mult_nth)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2207
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2208
  ultimately show ?thesis by blast
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2209
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2210
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2211
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2212
lemma radical_mult_distrib:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2213
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2214
  assumes
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2215
  ra0: "r k (a $ 0) ^ k = a $ 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2216
  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
  2217
  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
  2218
  and a0: "a$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2219
  and b0: "b$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2220
  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
  2221
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2222
  from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2223
    by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2224
  {assume "k=0" hence ?thesis by simp}
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2225
  moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2226
  {fix h assume k: "k = Suc h"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2227
  let ?ra = "fps_radical r (Suc h) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2228
  let ?rb = "fps_radical r (Suc h) b"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2229
  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
  2230
    using r0' k by (simp add: fps_mult_nth)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2231
  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
  2232
  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
  2233
    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
  2234
  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
  2235
ultimately show ?thesis by (cases k, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2236
qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2237
*)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2238
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2239
lemma fps_divide_1[simp]: "(a:: ('a::field) fps) / 1 = a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2240
  by (simp add: fps_divide_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2241
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2242
lemma radical_divide:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2243
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2244
  assumes kp: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2245
    and ra0: "(r k (a $ 0)) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2246
    and rb0: "(r k (b $ 0)) ^ k = b $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2247
    and a0: "a$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2248
    and b0: "b$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2249
  shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \<longleftrightarrow>
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2250
    fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2251
  (is "?lhs = ?rhs")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2252
proof -
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2253
  let ?r = "fps_radical r k"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2254
  from kp obtain h where k: "k = Suc h" by (cases k) auto
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2255
  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
  2256
  have rb0': "r k (b$0) \<noteq> 0" using b0 rb0 k by auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2257
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2258
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2259
    assume ?rhs
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2260
    then have "?r (a/b) $ 0 = (?r a / ?r b)$0" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2261
    then have ?lhs using k a0 b0 rb0'
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2262
      by (simp add: fps_divide_def fps_mult_nth fps_inverse_def divide_inverse)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2263
  }
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2264
  moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2265
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2266
    assume h: ?lhs
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2267
    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
  2268
      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
  2269
    have th0: "r k ((a/b)$0) ^ k = (a/b)$0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2270
      by (simp add: h nonzero_power_divide[OF rb0'] ra0 rb0)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2271
    from a0 b0 ra0' rb0' kp h
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2272
    have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2273
      by (simp add: fps_divide_def 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
  2274
    from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \<noteq> 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2275
      by (simp add: fps_divide_def fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero)
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2276
    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
  2277
    note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]]
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2278
    have th2: "(?r a / ?r b)^k = a/b"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2279
      by (simp add: fps_divide_def power_mult_distrib fps_inverse_power[symmetric])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2280
    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]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2281
    have ?rhs .
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2282
  }
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2283
  ultimately show ?thesis by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2284
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2285
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2286
lemma radical_inverse:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2287
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2288
  assumes k: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2289
    and ra0: "r k (a $ 0) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2290
    and r1: "(r k 1)^k = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2291
    and a0: "a$0 \<noteq> 0"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2292
  shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \<longleftrightarrow> fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2293
  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
  2294
  by (simp add: divide_inverse fps_divide_def)
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2295
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  2296
subsection{* Derivative of composition *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2297
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2298
lemma fps_compose_deriv:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2299
  fixes a:: "('a::idom) fps"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2300
  assumes b0: "b$0 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2301
  shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * (fps_deriv b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2302
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2303
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2304
    fix n
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2305
    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
  2306
      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
  2307
    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
  2308
      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
  2309
    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
  2310
      unfolding fps_mult_left_const_nth  by (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2311
    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
  2312
      unfolding fps_mult_nth ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2313
    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}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2314
      apply (rule setsum_mono_zero_right)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2315
      apply (auto simp add: mult_delta_left setsum_delta not_le)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2316
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2317
    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
  2318
      unfolding fps_deriv_nth
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2319
      apply (rule setsum_reindex_cong [where f = Suc])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2320
      by (auto simp add: mult_assoc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2321
    finally have th0: "(fps_deriv (a oo b))$n =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2322
      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
  2323
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2324
    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}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2325
      unfolding fps_mult_nth by (simp add: mult_ac)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2326
    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}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2327
      unfolding fps_deriv_nth fps_compose_nth setsum_right_distrib mult_assoc
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2328
      apply (rule setsum_cong2)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2329
      apply (rule setsum_mono_zero_left)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2330
      apply (simp_all add: subset_eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2331
      apply clarify
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2332
      apply (subgoal_tac "b^i$x = 0")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2333
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2334
      apply (rule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2335
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2336
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2337
    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
  2338
      unfolding setsum_right_distrib
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2339
      apply (subst setsum_commute)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2340
      apply (rule setsum_cong2)+
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2341
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2342
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2343
    finally have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2344
      unfolding th0 by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2345
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2346
  then show ?thesis by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2347
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2348
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2349
lemma fps_mult_X_plus_1_nth:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2350
  "((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
  2351
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2352
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2353
  then show ?thesis by (simp add: fps_mult_nth )
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2354
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2355
  case (Suc m)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2356
  have "((1+X)*a) $n = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0..n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2357
    by (simp add: fps_mult_nth)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2358
  also have "\<dots> = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2359
    unfolding Suc
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2360
    apply (rule setsum_mono_zero_right)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2361
    apply auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2362
    done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2363
  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
  2364
    by (simp add: Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2365
  finally show ?thesis .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2366
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2367
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  2368
subsection{* Finite FPS (i.e. polynomials) and X *}
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2369
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2370
lemma fps_poly_sum_X:
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2371
  assumes z: "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2372
  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
  2373
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2374
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2375
    fix i
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2376
    have "a$i = ?r$i"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2377
      unfolding fps_setsum_nth fps_mult_left_const_nth X_power_nth
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  2378
      by (simp add: mult_delta_right setsum_delta' z)
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  2379
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2380
  then show ?thesis unfolding fps_eq_iff by blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2381
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2382
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2383
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  2384
subsection{* Compositional inverses *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2385
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2386
fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::{field}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2387
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2388
  "compinv a 0 = X$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2389
| "compinv a (Suc n) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2390
    (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
  2391
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2392
definition "fps_inv a = Abs_fps (compinv a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2393
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2394
lemma fps_inv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2395
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2396
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2397
  shows "fps_inv a oo a = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2398
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2399
  let ?i = "fps_inv a oo a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2400
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2401
    fix n
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2402
    have "?i $n = X$n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2403
    proof (induct n rule: nat_less_induct)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2404
      fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2405
      assume h: "\<forall>m<n. ?i$m = X$m"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2406
      show "?i $ n = X$n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2407
      proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2408
        case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2409
        then show ?thesis using a0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2410
          by (simp add: fps_compose_nth fps_inv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2411
      next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2412
        case (Suc n1)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2413
        have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2414
          by (simp add: fps_compose_nth Suc startsby_zero_power_nth_same[OF a0] del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2415
        also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2416
          (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2417
          using a0 a1 Suc by (simp add: fps_inv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2418
        also have "\<dots> = X$n" using Suc by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2419
        finally show ?thesis .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2420
      qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2421
    qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2422
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2423
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2424
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2425
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2426
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2427
fun gcompinv :: "'a fps \<Rightarrow> 'a fps \<Rightarrow> nat \<Rightarrow> 'a::{field}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2428
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2429
  "gcompinv b a 0 = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2430
| "gcompinv b a (Suc n) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2431
    (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
  2432
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2433
definition "fps_ginv b a = Abs_fps (gcompinv b a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2434
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2435
lemma fps_ginv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2436
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2437
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2438
  shows "fps_ginv b a oo a = b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2439
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2440
  let ?i = "fps_ginv b a oo a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2441
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2442
    fix n
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2443
    have "?i $n = b$n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2444
    proof (induct n rule: nat_less_induct)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2445
      fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2446
      assume h: "\<forall>m<n. ?i$m = b$m"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2447
      show "?i $ n = b$n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2448
      proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2449
        case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2450
        then show ?thesis using a0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2451
          by (simp add: fps_compose_nth fps_ginv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2452
      next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2453
        case (Suc n1)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2454
        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"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2455
          by (simp add: fps_compose_nth Suc startsby_zero_power_nth_same[OF a0] del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2456
        also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2457
          (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2458
          using a0 a1 Suc by (simp add: fps_ginv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2459
        also have "\<dots> = b$n" using Suc by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2460
        finally show ?thesis .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2461
      qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2462
    qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2463
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2464
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2465
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2466
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2467
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
  2468
  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
  2469
  apply (induct_tac n rule: nat_less_induct)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2470
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2471
  apply (case_tac na)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2472
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2473
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2474
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2475
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2476
lemma fps_compose_1[simp]: "1 oo a = 1"
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30952
diff changeset
  2477
  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
  2478
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2479
lemma fps_compose_0[simp]: "0 oo a = 0"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  2480
  by (simp add: fps_eq_iff fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2481
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2482
lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a$0)"
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30952
diff changeset
  2483
  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2484
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2485
lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2486
  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_addf)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2487
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2488
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
  2489
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2490
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2491
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2492
  proof (rule finite_induct[OF True])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2493
    show "setsum f {} oo a = (\<Sum>i\<in>{}. f i oo a)" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2494
  next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2495
    fix x F
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2496
    assume fF: "finite F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2497
      and xF: "x \<notin> F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2498
      and h: "setsum f F oo a = setsum (\<lambda>i. f i oo a) F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2499
    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
  2500
      using fF xF h by (simp add: fps_compose_add_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2501
  qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2502
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2503
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2504
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2505
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2506
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2507
lemma convolution_eq:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2508
  "setsum (%i. a (i :: nat) * b (n - i)) {0 .. n} = setsum (%(i,j). a i * b j) {(i,j). i <= n \<and> j \<le> n \<and> i + j = n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2509
  apply (rule setsum_reindex_cong[where f=fst])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2510
  apply (clarsimp simp add: inj_on_def)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2511
  apply (auto simp add: set_eq_iff image_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2512
  apply (rule_tac x= "x" in exI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2513
  apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2514
  apply (rule_tac x="n - x" in exI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2515
  apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2516
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2517
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2518
lemma product_composition_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2519
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2520
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2521
  shows "((a oo c) * (b oo d))$n =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2522
    setsum (%(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}" (is "?l = ?r")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2523
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2524
  let ?S = "{(k\<Colon>nat, m\<Colon>nat). k + m \<le> n}"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2525
  have s: "?S \<subseteq> {0..n} <*> {0..n}" by (auto simp add: subset_eq)
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2526
  have f: "finite {(k\<Colon>nat, m\<Colon>nat). k + m \<le> n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2527
    apply (rule finite_subset[OF s])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2528
    apply auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2529
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2530
  have "?r =  setsum (%i. setsum (%(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2531
    apply (simp add: fps_mult_nth setsum_right_distrib)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2532
    apply (subst setsum_commute)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2533
    apply (rule setsum_cong2)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2534
    apply (auto simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2535
    done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2536
  also have "\<dots> = ?l"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2537
    apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2538
    apply (rule setsum_cong2)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2539
    apply (simp add: setsum_cartesian_product mult_assoc)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2540
    apply (rule setsum_mono_zero_right[OF f])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2541
    apply (simp add: subset_eq) apply presburger
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2542
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2543
    apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2544
    apply (clarsimp simp add: not_le)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2545
    apply (case_tac "x < aa")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2546
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2547
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2548
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2549
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2550
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2551
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2552
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2553
  finally show ?thesis by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2554
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2555
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2556
lemma product_composition_lemma':
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2557
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2558
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2559
  shows "((a oo c) * (b oo d))$n =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2560
    setsum (%k. setsum (%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
  2561
  unfolding product_composition_lemma[OF c0 d0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2562
  unfolding setsum_cartesian_product
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2563
  apply (rule setsum_mono_zero_left)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2564
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2565
  apply (clarsimp simp add: subset_eq)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2566
  apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2567
  apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2568
  apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2569
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2570
  unfolding fps_mult_nth
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2571
  apply (rule setsum_0')
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2572
  apply (clarsimp simp add: not_le)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  2573
  apply (case_tac "x < aa")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2574
  apply (rule startsby_zero_power_prefix[OF c0, rule_format])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2575
  apply simp
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  2576
  apply (subgoal_tac "n - x < ba")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2577
  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
  2578
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2579
  apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2580
  done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2581
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2582
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2583
lemma setsum_pair_less_iff:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2584
  "setsum (%((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2585
    setsum (%s. setsum (%i. a i * b (s - i) * c s) {0..s}) {0..n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2586
  (is "?l = ?r")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2587
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2588
  let ?KM = "{(k,m). k + m \<le> n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2589
  let ?f = "%s. UNION {(0::nat)..s} (%i. {(i,s - i)})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2590
  have th0: "?KM = UNION {0..n} ?f"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2591
    apply (simp add: set_eq_iff)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2592
    apply presburger (* FIXME: slow! *)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2593
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2594
  show "?l = ?r "
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2595
    unfolding th0
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2596
    apply (subst setsum_UN_disjoint)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2597
    apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2598
    apply (subst setsum_UN_disjoint)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2599
    apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2600
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2601
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2602
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2603
lemma fps_compose_mult_distrib_lemma:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2604
  assumes c0: "c$0 = (0::'a::idom)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2605
  shows "((a oo c) * (b oo c))$n =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2606
    setsum (%s. setsum (%i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2607
    (is "?l = ?r")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2608
  unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2609
  unfolding setsum_pair_less_iff[where a = "%k. a$k" and b="%m. b$m" and c="%s. (c ^ s)$n" and n = n] ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2610
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2611
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2612
lemma fps_compose_mult_distrib:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2613
  assumes c0: "c$0 = (0::'a::idom)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2614
  shows "(a * b) oo c = (a oo c) * (b oo c)" (is "?l = ?r")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2615
  apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma[OF c0])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2616
  apply (simp add: fps_compose_nth fps_mult_nth setsum_left_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2617
  done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2618
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2619
lemma fps_compose_setprod_distrib:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2620
  assumes c0: "c$0 = (0::'a::idom)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2621
  shows "(setprod a S) oo c = setprod (%k. a k oo c) S" (is "?l = ?r")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2622
  apply (cases "finite S")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2623
  apply simp_all
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2624
  apply (induct S rule: finite_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2625
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2626
  apply (simp add: fps_compose_mult_distrib[OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2627
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2628
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2629
lemma fps_compose_power:   assumes c0: "c$0 = (0::'a::idom)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2630
  shows "(a oo c)^n = a^n oo c" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2631
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2632
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2633
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2634
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2635
  case (Suc m)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2636
  have th0: "a^n = setprod (%k. a) {0..m}" "(a oo c) ^ n = setprod (%k. a oo c) {0..m}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2637
    by (simp_all add: setprod_constant Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2638
  then show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2639
    by (simp add: fps_compose_setprod_distrib[OF c0])
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2640
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2641
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2642
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
  2643
  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
  2644
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2645
lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2646
  unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2647
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2648
lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2649
  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
  2650
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2651
lemma fps_inverse_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2652
  assumes b0: "(b$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2653
    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
  2654
  shows "inverse a oo b = inverse (a oo b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2655
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2656
  let ?ia = "inverse a"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2657
  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
  2658
  let ?iab = "inverse ?ab"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2659
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2660
  from a0 have ia0: "?ia $ 0 \<noteq> 0" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2661
  from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2662
  have "(?ia oo b) *  (a oo b) = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2663
    unfolding fps_compose_mult_distrib[OF b0, symmetric]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2664
    unfolding inverse_mult_eq_1[OF a0]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2665
    fps_compose_1 ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2666
  
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2667
  then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2668
  then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2669
  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
  2670
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2671
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2672
lemma fps_divide_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2673
  assumes c0: "(c$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2674
    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
  2675
  shows "(a/b) oo c = (a oo c) / (b oo c)"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2676
    unfolding fps_divide_def fps_compose_mult_distrib[OF c0]
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2677
    fps_inverse_compose[OF c0 b0] ..
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2678
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2679
lemma gp:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2680
  assumes a0: "a$0 = (0::'a::field)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2681
  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2682
    (is "?one oo a = _")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2683
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2684
  have o0: "?one $ 0 \<noteq> 0" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2685
  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
  2686
  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
  2687
  have "inverse ?one = 1 - X" by (simp add: fps_eq_iff)
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2688
  hence "inverse (inverse ?one) = inverse (1 - X)" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2689
  hence 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
  2690
    by (simp add: fps_divide_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2691
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2692
    unfolding th
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2693
    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
  2694
    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
  2695
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2696
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2697
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
  2698
  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
  2699
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2700
lemma fps_compose_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2701
  assumes b0: "b$0 = (0::'a::field_char_0)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2702
    and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2703
    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
  2704
  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
  2705
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2706
  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
  2707
  let ?ab = "a oo b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2708
  have ab0: "?ab $ 0 = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2709
    by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2710
  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
  2711
    by simp_all
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2712
  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
  2713
    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
  2714
  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
  2715
    unfolding fps_compose_power[OF b0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2716
    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  ..
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2717
  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
  2718
  show ?thesis  .
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2719
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  2720
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2721
lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2722
  by (simp add: fps_eq_iff fps_compose_nth setsum_right_distrib mult_assoc)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2723
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2724
lemma fps_const_mult_apply_right:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2725
  "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2726
  by (auto simp add: fps_const_mult_apply_left mult_commute)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2727
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2728
lemma fps_compose_assoc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2729
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2730
    and b0: "b$0 = 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2731
  shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2732
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2733
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2734
    fix n
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2735
    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
  2736
      by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2737
        setsum_right_distrib mult_assoc fps_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2738
    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
  2739
      by (simp add: fps_compose_setsum_distrib)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2740
    also have "\<dots> = ?r$n"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2741
      apply (simp add: fps_compose_nth fps_setsum_nth setsum_left_distrib mult_assoc)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2742
      apply (rule setsum_cong2)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2743
      apply (rule setsum_mono_zero_right)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2744
      apply (auto simp add: not_le)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2745
      apply (erule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2746
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2747
    finally have "?l$n = ?r$n" .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2748
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2749
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2750
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2751
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2752
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2753
lemma fps_X_power_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2754
  assumes a0: "a$0=0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2755
  shows "X^k oo a = (a::('a::idom fps))^k" (is "?l = ?r")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2756
proof (cases k)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2757
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2758
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2759
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2760
  case(Suc h)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2761
  {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2762
    fix n
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2763
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2764
      assume kn: "k>n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2765
      hence "?l $ n = ?r $n" using a0 startsby_zero_power_prefix[OF a0] Suc
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2766
        by (simp add: fps_compose_nth del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2767
    }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2768
    moreover
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2769
    {
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2770
      assume kn: "k \<le> n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2771
      hence "?l$n = ?r$n"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2772
        by (simp add: fps_compose_nth mult_delta_left setsum_delta)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2773
    }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2774
    moreover have "k >n \<or> k\<le> n"  by arith
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2775
    ultimately have "?l$n = ?r$n"  by blast
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2776
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2777
  then show ?thesis unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2778
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2779
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2780
lemma fps_inv_right:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2781
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2782
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2783
  shows "a oo fps_inv a = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2784
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2785
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2786
  let ?iaa = "a oo fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2787
  have th0: "?ia $ 0 = 0" by (simp add: fps_inv_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2788
  have th1: "?iaa $ 0 = 0" using a0 a1
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2789
    by (simp add: fps_inv_def fps_compose_nth)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2790
  have th2: "X$0 = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2791
  from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo X" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2792
  then have "(a oo fps_inv a) oo a = X oo a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2793
    by (simp add: fps_compose_assoc[OF a0 th0] X_fps_compose_startby0[OF a0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2794
  with fps_compose_inj_right[OF a0 a1]
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2795
  show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2796
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2797
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2798
lemma fps_inv_deriv:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2799
  assumes a0:"a$0 = (0::'a::{field})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2800
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2801
  shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2802
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2803
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2804
  let ?d = "fps_deriv a oo ?ia"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2805
  let ?dia = "fps_deriv ?ia"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2806
  have ia0: "?ia$0 = 0" by (simp add: fps_inv_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2807
  have th0: "?d$0 \<noteq> 0" using a1 by (simp add: fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2808
  from fps_inv_right[OF a0 a1] have "?d * ?dia = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2809
    by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2810
  hence "inverse ?d * ?d * ?dia = inverse ?d * 1" by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2811
  with inverse_mult_eq_1 [OF th0]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2812
  show "?dia = inverse ?d" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2813
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2814
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2815
lemma fps_inv_idempotent:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2816
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2817
    and a1: "a$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2818
  shows "fps_inv (fps_inv a) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2819
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2820
  let ?r = "fps_inv"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2821
  have ra0: "?r a $ 0 = 0" by (simp add: fps_inv_def)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2822
  from a1 have ra1: "?r a $ 1 \<noteq> 0" by (simp add: fps_inv_def field_simps)
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2823
  have X0: "X$0 = 0" by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2824
  from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2825
  then have "?r (?r a) oo ?r a oo a = X oo a" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2826
  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
  2827
    unfolding X_fps_compose_startby0[OF a0]
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2828
    unfolding fps_compose_assoc[OF a0 ra0, symmetric] .
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2829
  then show ?thesis unfolding fps_inv[OF a0 a1] by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2830
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2831
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2832
lemma fps_ginv_ginv:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2833
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2834
    and a1: "a$1 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2835
    and c0: "c$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2836
    and  c1: "c$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2837
  shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2838
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2839
  let ?r = "fps_ginv"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2840
  from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2841
  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0" by (simp add: fps_ginv_def field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2842
  from fps_ginv[OF rca0 rca1]
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2843
  have "?r b (?r c a) oo ?r c a = b" .
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2844
  then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2845
  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
  2846
    apply (subst fps_compose_assoc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2847
    using a0 c0 apply (auto simp add: fps_ginv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2848
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2849
  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
  2850
    unfolding fps_ginv[OF a0 a1] .
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2851
  then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c" by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2852
  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
  2853
    apply (subst fps_compose_assoc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2854
    using a0 c0 apply (auto simp add: fps_inv_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2855
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2856
  then show ?thesis unfolding fps_inv_right[OF c0 c1] by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2857
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2858
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2859
lemma fps_ginv_deriv:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2860
  assumes a0:"a$0 = (0::'a::{field})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2861
    and a1: "a$1 \<noteq> 0"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2862
  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
  2863
proof -
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2864
  let ?ia = "fps_ginv b a"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2865
  let ?iXa = "fps_ginv X a"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2866
  let ?d = "fps_deriv"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2867
  let ?dia = "?d ?ia"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2868
  have iXa0: "?iXa $ 0 = 0" by (simp add: fps_ginv_def)
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2869
  have da0: "?d a $ 0 \<noteq> 0" using a1 by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2870
  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b" by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2871
  then have "(?d ?ia oo a) * ?d a = ?d b" unfolding fps_compose_deriv[OF a0] .
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2872
  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2873
  then have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2874
    by (simp add: fps_divide_def)
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2875
  then have "(?d ?ia oo a) oo ?iXa =  (?d b / ?d a) oo ?iXa "
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2876
    unfolding inverse_mult_eq_1[OF da0] by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2877
  then have "?d ?ia oo (a oo ?iXa) =  (?d b / ?d a) oo ?iXa"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2878
    unfolding fps_compose_assoc[OF iXa0 a0] .
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2879
  then show ?thesis unfolding fps_inv_ginv[symmetric]
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2880
    unfolding fps_inv_right[OF a0 a1] by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2881
qed
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  2882
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  2883
subsection{* Elementary series *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2884
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  2885
subsubsection{* Exponential series *}
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2886
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
  2887
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2888
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
  2889
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2890
  { fix n
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2891
    have "?l$n = ?r $ n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2892
      apply (auto simp add: E_def field_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2893
      apply (simp add: of_nat_mult field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2894
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2895
  }
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2896
  then show ?thesis by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2897
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2898
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2899
lemma E_unique_ODE:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2900
  "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
  2901
  (is "?lhs \<longleftrightarrow> ?rhs")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2902
proof
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2903
  assume d: ?lhs
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2904
  from d have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2905
    by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2906
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2907
    fix n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2908
    have "a$n = a$0 * c ^ n/ (of_nat (fact n))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2909
      apply (induct n)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2910
      apply simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2911
      unfolding th
32042
df28ead1cf19 Repairs regarding new Fact.thy.
avigad
parents: 31968
diff changeset
  2912
      using fact_gt_zero_nat
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2913
      apply (simp add: field_simps del: of_nat_Suc fact_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2914
      apply (drule sym)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2915
      apply (simp add: field_simps of_nat_mult)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2916
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2917
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2918
  note th' = this
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2919
  show ?rhs by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro: th')
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2920
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2921
  assume h: ?rhs
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2922
  show ?lhs
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2923
    apply (subst h)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2924
    apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2925
    apply (simp only: h[symmetric])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2926
    apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2927
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2928
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2929
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2930
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
  2931
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2932
  have "fps_deriv (?r) = fps_const (a+b) * ?r"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2933
    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2934
  then have "?r = ?l" apply (simp only: E_unique_ODE)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2935
    by (simp add: fps_mult_nth E_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2936
  then show ?thesis ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2937
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2938
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2939
lemma E_nth[simp]: "E a $ n = a^n / of_nat (fact n)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2940
  by (simp add: E_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2941
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  2942
lemma E0[simp]: "E (0::'a::{field}) = 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2943
  by (simp add: fps_eq_iff power_0_left)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2944
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2945
lemma E_neg: "E (- a) = inverse (E (a::'a::field_char_0))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2946
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2947
  from E_add_mult[of a "- a"] have th0: "E a * E (- a) = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2948
    by (simp )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2949
  have th1: "E a $ 0 \<noteq> 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2950
  from fps_inverse_unique[OF th1 th0] show ?thesis by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2951
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2952
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2953
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
  2954
  by (induct n) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2955
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  2956
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
  2957
  by (simp add: fps_eq_iff X_fps_compose)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2958
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2959
lemma LE_compose:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2960
  assumes a: "a\<noteq>0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2961
  shows "fps_inv (E a - 1) oo (E a - 1) = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2962
    and "(E a - 1) oo fps_inv (E a - 1) = X"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2963
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2964
  let ?b = "E a - 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2965
  have b0: "?b $ 0 = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2966
  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2967
  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
  2968
  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
  2969
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2970
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2971
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2972
lemma fps_const_inverse:
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2973
  "a \<noteq> 0 \<Longrightarrow> inverse (fps_const (a::'a::field)) = fps_const (inverse a)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2974
  apply (auto simp add: fps_eq_iff fps_inverse_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2975
  apply (case_tac n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2976
  apply auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2977
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2978
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2979
lemma inverse_one_plus_X:
31021
53642251a04f farewell to class recpower
haftmann
parents: 30994
diff changeset
  2980
  "inverse (1 + X) = Abs_fps (\<lambda>n. (- 1 ::'a::{field})^n)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2981
  (is "inverse ?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2982
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2983
  have th: "?l * ?r = 1"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  2984
    by (auto simp add: field_simps fps_eq_iff minus_one_power_iff simp del: minus_one)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2985
  have th': "?l $ 0 \<noteq> 0" by (simp add: )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2986
  from fps_inverse_unique[OF th' th] show ?thesis .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2987
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2988
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2989
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
  2990
  by (induct n) (auto simp add: field_simps E_add_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2991
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  2992
lemma radical_E:
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2993
  assumes r: "r (Suc k) 1 = 1"
31370
chaieb
parents: 31274 31369
diff changeset
  2994
  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
  2995
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2996
  let ?ck = "(c / of_nat (Suc k))"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2997
  let ?r = "fps_radical r (Suc k)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  2998
  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
  2999
    by (simp_all del: of_nat_Suc)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3000
  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
  3001
  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
  3002
    "r (Suc k) (E c $ 0) = E ?ck $ 0" "E c $ 0 \<noteq> 0" using r by simp_all
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3003
  from th0 radical_unique[where r=r and k=k, OF th]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3004
  show ?thesis by auto
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3005
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3006
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3007
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
  3008
  apply (auto simp add: fps_eq_iff E_def fps_compose_def power_mult_distrib)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3009
  apply (simp add: cond_value_iff cond_application_beta setsum_delta' cong del: if_weak_cong)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3010
  done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3011
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3012
text{* The generalized binomial theorem as a  consequence of @{thm E_add_mult} *}
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3013
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3014
lemma gbinomial_theorem:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3015
  "((a::'a::{field_char_0, field_inverse_zero})+b) ^ n =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3016
    (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3017
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3018
  from E_add_mult[of a b]
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3019
  have "(E (a + b)) $ n = (E a * E b)$n" by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3020
  then have "(a + b) ^ n =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3021
    (\<Sum>i\<Colon>nat = 0\<Colon>nat..n. a ^ i * b ^ (n - i)  * (of_nat (fact n) / of_nat (fact i * fact (n - i))))"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3022
    by (simp add: field_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3023
  then show ?thesis
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3024
    apply simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3025
    apply (rule setsum_cong2)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3026
    apply simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3027
    apply (frule binomial_fact[where ?'a = 'a, symmetric])
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3028
    by (simp add: field_simps of_nat_mult)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3029
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3030
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3031
text{* And the nat-form -- also available from Binomial.thy *}
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3032
lemma binomial_theorem: "(a+b) ^ n = (\<Sum>k=0..n. (n choose k) * a^k * b^(n-k))"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3033
  using gbinomial_theorem[of "of_nat a" "of_nat b" n]
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3034
  unfolding of_nat_add[symmetric] of_nat_power[symmetric] of_nat_mult[symmetric] of_nat_setsum[symmetric]
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3035
  by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3036
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3037
subsubsection{* Logarithmic series *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3038
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3039
lemma Abs_fps_if_0:
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3040
  "Abs_fps(%n. if n=0 then (v::'a::ring_1) else f n) = fps_const v + X * Abs_fps (%n. f (Suc n))"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3041
  by (auto simp add: fps_eq_iff)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3042
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3043
definition L :: "'a::field_char_0 \<Rightarrow> 'a fps"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3044
  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
  3045
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3046
lemma fps_deriv_L: "fps_deriv (L c) = fps_const (1/c) * inverse (1 + X)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3047
  unfolding inverse_one_plus_X
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3048
  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
  3049
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3050
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
  3051
  by (simp add: L_def field_simps)
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3052
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3053
lemma L_0[simp]: "L c $ 0 = 0" by (simp add: L_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3054
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3055
lemma L_E_inv:
31370
chaieb
parents: 31274 31369
diff changeset
  3056
  assumes a: "a\<noteq> (0::'a::{field_char_0})"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3057
  shows "L a = fps_inv (E a - 1)" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3058
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3059
  let ?b = "E a - 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3060
  have b0: "?b $ 0 = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3061
  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3062
  have "fps_deriv (E a - 1) oo fps_inv (E a - 1) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3063
    (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
  3064
    by (simp add: field_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3065
  also have "\<dots> = fps_const a * (X + 1)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3066
    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
  3067
    apply (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3068
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3069
  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
  3070
  from fps_inv_deriv[OF b0 b1, unfolded eq]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3071
  have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3072
    using a
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3073
    by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3074
  hence "fps_deriv ?l = fps_deriv ?r"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3075
    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
  3076
  then show ?thesis unfolding fps_deriv_eq_iff
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3077
    by (simp add: L_nth fps_inv_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3078
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3079
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3080
lemma L_mult_add:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3081
  assumes c0: "c\<noteq>0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3082
    and d0: "d\<noteq>0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3083
  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
  3084
  (is "?r = ?l")
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3085
proof-
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3086
  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
  3087
  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
  3088
    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
  3089
  also have "\<dots> = fps_deriv ?l"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3090
    apply (simp add: fps_deriv_L)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3091
    apply (simp add: fps_eq_iff eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3092
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3093
  finally show ?thesis
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3094
    unfolding fps_deriv_eq_iff by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3095
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3096
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3097
subsubsection{* Binomial series *}
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3098
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3099
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
  3100
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3101
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
  3102
  by (simp add: fps_binomial_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3103
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3104
lemma fps_binomial_ODE_unique:
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3105
  fixes c :: "'a::field_char_0"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3106
  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
  3107
  (is "?lhs \<longleftrightarrow> ?rhs")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3108
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3109
  let ?da = "fps_deriv a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3110
  let ?x1 = "(1 + X):: 'a fps"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3111
  let ?l = "?x1 * ?da"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3112
  let ?r = "fps_const c * a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3113
  have x10: "?x1 $ 0 \<noteq> 0" by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3114
  have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3115
  also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3116
    apply (simp only: fps_divide_def  mult_assoc[symmetric] inverse_mult_eq_1[OF x10])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3117
    apply (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3118
    done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3119
  finally have eq: "?l = ?r \<longleftrightarrow> ?lhs" by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3120
  moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3121
  {assume h: "?l = ?r"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3122
    {fix n
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3123
      from h have lrn: "?l $ n = ?r$n" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3124
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3125
      from lrn
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3126
      have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3127
        apply (simp add: field_simps del: of_nat_Suc)
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3128
        by (cases n, simp_all add: field_simps del: of_nat_Suc)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3129
    }
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3130
    note th0 = this
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3131
    {fix n have "a$n = (c gchoose n) * a$0"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3132
      proof(induct n)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3133
        case 0 thus ?case by simp
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3134
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3135
        case (Suc m)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3136
        thus ?case unfolding th0
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3137
          apply (simp add: field_simps del: of_nat_Suc)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3138
          unfolding mult_assoc[symmetric] gbinomial_mult_1
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3139
          by (simp add: field_simps)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3140
      qed}
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3141
    note th1 = this
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3142
    have ?rhs
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3143
      apply (simp add: fps_eq_iff)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3144
      apply (subst th1)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3145
      by (simp add: field_simps)}
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3146
  moreover
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3147
  {assume h: ?rhs
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3148
  have th00:"\<And>x y. x * (a$0 * y) = a$0 * (x*y)" by (simp add: mult_commute)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3149
    have "?l = ?r"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3150
      apply (subst h)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3151
      apply (subst (2) h)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3152
      apply (clarsimp simp add: fps_eq_iff field_simps)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3153
      unfolding mult_assoc[symmetric] th00 gbinomial_mult_1
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3154
      by (simp add: field_simps gbinomial_mult_1)}
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3155
  ultimately show ?thesis by blast
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3156
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3158
lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + X)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3159
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3160
  let ?a = "fps_binomial c"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3161
  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
  3162
  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
  3163
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3164
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3165
lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r")
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3166
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3167
  let ?P = "?r - ?l"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3168
  let ?b = "fps_binomial"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3169
  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
  3170
  have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3171
  also have "\<dots> = inverse (1 + X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3172
    unfolding fps_binomial_deriv
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3173
    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
  3174
  also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3175
    by (simp add: field_simps fps_divide_def 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
  3176
  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
  3177
    by (simp add: fps_divide_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3178
  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
  3179
    unfolding fps_binomial_ODE_unique[symmetric]
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3180
    using th0 by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3181
  hence "?P = 0" by (simp add: fps_mult_nth)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3182
  then show ?thesis by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3183
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3184
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3185
lemma fps_minomial_minus_one: "fps_binomial (- 1) = inverse (1 + X)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3186
  (is "?l = inverse ?r")
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3187
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3188
  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
  3189
  have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  3190
    by (simp add: fps_inverse_deriv[OF th] fps_divide_def power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg minus_one)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3191
  have eq: "inverse ?r $ 0 = 1"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3192
    by (simp add: fps_inverse_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3193
  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
  3194
  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
  3195
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3196
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3197
text{* Vandermonde's Identity as a consequence *}
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3198
lemma gbinomial_Vandermonde: "setsum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3199
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3200
  let ?ba = "fps_binomial a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3201
  let ?bb = "fps_binomial b"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3202
  let ?bab = "fps_binomial (a + b)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3203
  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
  3204
  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
  3205
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3206
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3207
lemma binomial_Vandermonde: "setsum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3208
  using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3209
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3210
  apply (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric] of_nat_setsum[symmetric] of_nat_add[symmetric])
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3211
  by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3212
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  3213
lemma binomial_Vandermonde_same: "setsum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2*n) choose n"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3214
  using binomial_Vandermonde[of n n n,symmetric]
47217
501b9bbd0d6e removed redundant nat-specific copies of theorems
huffman
parents: 47108
diff changeset
  3215
  unfolding mult_2 apply (simp add: power2_eq_square)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3216
  apply (rule setsum_cong2)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3217
  by (auto intro:  binomial_symmetric)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3218
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3219
lemma Vandermonde_pochhammer_lemma:
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3220
  fixes a :: "'a::field_char_0"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3221
  assumes b: "\<forall> j\<in>{0 ..<n}. b \<noteq> of_nat j"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3222
  shows "setsum (%k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} = pochhammer (- (a+ b)) n / pochhammer (- b) n" (is "?l = ?r")
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3223
proof-
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3224
  let ?m1 = "%m. (- 1 :: 'a) ^ m"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3225
  let ?f = "%m. of_nat (fact m)"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3226
  let ?p = "%(x::'a). pochhammer (- x)"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3227
  from b have bn0: "?p b n \<noteq> 0" unfolding pochhammer_eq_0_iff by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3228
  {fix k assume kn: "k \<in> {0..n}"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3229
    {assume c:"pochhammer (b - of_nat n + 1) n = 0"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3230
      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
  3231
        unfolding pochhammer_eq_0_iff by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3232
      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
  3233
        by (simp add: algebra_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3234
      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
  3235
        using j kn by (simp add: of_nat_diff)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3236
      with b have False using j by auto}
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3237
    then have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3238
      by (auto simp add: algebra_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3239
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3240
    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
  3241
      by (rule pochhammer_neq_0_mono)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3242
    {assume k0: "k = 0 \<or> n =0"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3243
      then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f 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
  3244
        using kn
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3245
        by (cases "k=0", simp_all add: gbinomial_pochhammer)}
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3246
    moreover
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3247
    {assume n0: "n \<noteq> 0" and k0: "k \<noteq> 0"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3248
      then obtain m where m: "n = Suc m" by (cases n, auto)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3249
      from k0 obtain h where h: "k = Suc h" by (cases k, auto)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3250
      {assume kn: "k = n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3251
        then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3252
          using kn pochhammer_minus'[where k=k and n=n and b=b]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3253
          apply (simp add:  pochhammer_same)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3254
          using bn0
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3255
          by (simp add: field_simps power_add[symmetric])}
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3256
      moreover
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3257
      {assume nk: "k \<noteq> n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3258
        have m1nk: "?m1 n = setprod (%i. - 1) {0..m}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3259
          "?m1 k = setprod (%i. - 1) {0..h}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3260
          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
  3261
        from kn nk have kn': "k < n" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3262
        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3263
          using bn0 kn
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3264
          unfolding pochhammer_eq_0_iff
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3265
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3266
          apply (erule_tac x= "n - ka - 1" in allE)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3267
          by (auto simp add: algebra_simps of_nat_diff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3268
        have eq1: "setprod (%k. (1::'a) + of_nat m - of_nat k) {0 .. h} = 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
  3269
          apply (rule strong_setprod_reindex_cong[where f="%k. Suc m - k "])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3270
          using kn' h m
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3271
          apply (auto simp add: inj_on_def image_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3272
          apply (rule_tac x="Suc m - x" in bexI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3273
          apply (simp_all add: of_nat_diff)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3274
          done
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3275
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3276
        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
  3277
          unfolding m1nk
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3278
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3279
          unfolding m h pochhammer_Suc_setprod
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3280
          apply (simp add: field_simps del: fact_Suc minus_one)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3281
          unfolding fact_altdef_nat id_def
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3282
          unfolding of_nat_setprod
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3283
          unfolding setprod_timesf[symmetric]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3284
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3285
          unfolding eq1
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3286
          apply (subst setprod_Un_disjoint[symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3287
          apply (auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3288
          apply (rule setprod_cong)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3289
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3290
          done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3291
        have th20: "?m1 n * ?p b n = setprod (%i. b - of_nat i) {0..m}"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3292
          unfolding m1nk
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3293
          unfolding m h pochhammer_Suc_setprod
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3294
          unfolding setprod_timesf[symmetric]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3295
          apply (rule setprod_cong)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3296
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3297
          done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3298
        have th21:"pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {n - k .. n - 1}"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3299
          unfolding h m
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3300
          unfolding pochhammer_Suc_setprod
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3301
          apply (rule strong_setprod_reindex_cong[where f="%k. n - 1 - k"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3302
          using kn
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3303
          apply (auto simp add: inj_on_def m h image_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3304
          apply (rule_tac x= "m - x" in bexI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3305
          by (auto simp add: of_nat_diff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3306
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3307
        have "?m1 n * ?p b n = pochhammer (b - of_nat n + 1) k * setprod (%i. b - of_nat i) {0.. n - k - 1}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3308
          unfolding th20 th21
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3309
          unfolding h m
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3310
          apply (subst setprod_Un_disjoint[symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3311
          using kn' h m
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3312
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3313
          apply (rule setprod_cong)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3314
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3315
          done
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3316
        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = setprod (%i. b - of_nat i) {0.. n - k - 1}"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3317
          using nz' by (simp add: field_simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3318
        have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3319
          using bnz0
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3320
          by (simp add: field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3321
        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
  3322
          unfolding th1 th2
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3323
          using kn' by (simp add: gbinomial_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3324
        finally have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" by simp}
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3325
      ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f 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
  3326
        by (cases "k =n", auto)}
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3327
    ultimately have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" "pochhammer (1 + b - of_nat n) k \<noteq> 0 "
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3328
      using nz'
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3329
      apply (cases "n=0", auto)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3330
      by (cases "k", auto)}
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3331
  note th00 = this
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3332
  have "?r = ((a + b) gchoose n) * (of_nat (fact n)/ (?m1 n * pochhammer (- b) n))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3333
    unfolding gbinomial_pochhammer
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3334
    using bn0 by (auto simp add: field_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3335
  also have "\<dots> = ?l"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3336
    unfolding gbinomial_Vandermonde[symmetric]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3337
    apply (simp add: th00)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3338
    unfolding gbinomial_pochhammer
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3339
    using bn0 apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3340
    apply (rule setsum_cong2)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3341
    apply (drule th00(2))
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3342
    by (simp add: field_simps power_add[symmetric])
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3343
  finally show ?thesis by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3344
qed
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3345
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3346
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3347
lemma Vandermonde_pochhammer:
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3348
   fixes a :: "'a::field_char_0"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3349
  assumes c: "ALL i : {0..< n}. c \<noteq> - of_nat i"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3350
  shows "setsum (%k. (pochhammer a k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3351
proof-
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3352
  let ?a = "- a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3353
  let ?b = "c + of_nat n - 1"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3354
  have h: "\<forall> j \<in>{0..< n}. ?b \<noteq> of_nat j" using c
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3355
    apply (auto simp add: algebra_simps of_nat_diff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3356
    apply (erule_tac x= "n - j - 1" in ballE)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3357
    by (auto simp add: of_nat_diff algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3358
  have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3359
    unfolding pochhammer_minus[OF le_refl]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3360
    by (simp add: algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3361
  have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3362
    unfolding pochhammer_minus[OF le_refl]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3363
    by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3364
  have nz: "pochhammer c n \<noteq> 0" using c
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3365
    by (simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3366
  from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3367
  show ?thesis 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
  3368
qed
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3369
29906
80369da39838 section -> subsection
huffman
parents: 29692
diff changeset
  3370
subsubsection{* Formal trigonometric functions  *}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3371
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3372
definition "fps_sin (c::'a::field_char_0) =
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3373
  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
  3374
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3375
definition "fps_cos (c::'a::field_char_0) =
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3376
  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
  3377
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3378
lemma fps_sin_deriv:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3379
  "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3380
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3381
proof (rule fps_ext)
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3382
  fix n::nat
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3383
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3384
    assume en: "even n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3385
    have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3386
    also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3387
      using en by (simp add: fps_sin_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3388
    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
  3389
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3390
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3391
    also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3392
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3393
    finally have "?lhs $n = ?rhs$n" using en
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3394
      by (simp add: fps_cos_def field_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3395
  }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3396
  then show "?lhs $ n = ?rhs $ n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3397
    by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3398
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3399
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3400
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
  3401
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3402
proof (rule fps_ext)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3403
  have th0: "\<And>n. - ((- 1::'a) ^ n) = (- 1)^Suc n" by simp
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3404
  have th1: "\<And>n. odd n \<Longrightarrow> Suc ((n - 1) div 2) = Suc n div 2"
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3405
    by (case_tac n, simp_all)
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3406
  fix n::nat
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3407
  {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3408
    assume en: "odd n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3409
    from en have n0: "n \<noteq>0 " by presburger
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3410
    have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3411
    also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3412
      using en by (simp add: fps_cos_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3413
    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
  3414
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3415
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3416
    also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3417
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3418
    also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3419
      unfolding th0 unfolding th1[OF en] by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3420
    finally have "?lhs $n = ?rhs$n" using en
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3421
      by (simp add: fps_sin_def field_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3422
  }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3423
  then show "?lhs $ n = ?rhs $ n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3424
    by (cases "even n") (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3425
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3426
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3427
lemma fps_sin_cos_sum_of_squares:
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  3428
  "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1" (is "?lhs = 1")
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  3429
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3430
  have "fps_deriv ?lhs = 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3431
    apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3432
    apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3433
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3434
  then have "?lhs = fps_const (?lhs $ 0)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3435
    unfolding fps_deriv_eq_0_iff .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3436
  also have "\<dots> = 1"
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30952
diff changeset
  3437
    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
  3438
  finally show ?thesis .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3439
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3440
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3441
lemma divide_eq_iff: "a \<noteq> (0::'a::field) \<Longrightarrow> x / a = y \<longleftrightarrow> x = y * a"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3442
by auto
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3443
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3444
lemma eq_divide_iff: "a \<noteq> (0::'a::field) \<Longrightarrow> x = y / a \<longleftrightarrow> x * a = y"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3445
by auto
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3446
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3447
lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3448
unfolding fps_sin_def by simp
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3449
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3450
lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3451
unfolding fps_sin_def by simp
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3452
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3453
lemma fps_sin_nth_add_2:
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3454
  "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat(n+1) * of_nat(n+2)))"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3455
unfolding fps_sin_def
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3456
apply (cases n, simp)
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3457
apply (simp add: divide_eq_iff eq_divide_iff del: of_nat_Suc fact_Suc)
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3458
apply (simp add: of_nat_mult del: of_nat_Suc mult_Suc)
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3459
done
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3460
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3461
lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3462
unfolding fps_cos_def by simp
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3463
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3464
lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0"
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3465
unfolding fps_cos_def by simp
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3466
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3467
lemma fps_cos_nth_add_2:
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3468
  "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
  3469
  unfolding fps_cos_def
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3470
  apply (simp add: divide_eq_iff eq_divide_iff del: of_nat_Suc fact_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3471
  apply (simp add: of_nat_mult del: of_nat_Suc mult_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3472
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3473
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3474
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
  3475
  unfolding One_nat_def numeral_2_eq_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3476
  apply (induct n rule: nat_less_induct)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3477
  apply (case_tac n, simp)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3478
  apply (rename_tac m, case_tac m, simp)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3479
  apply (rename_tac k, case_tac k, simp_all)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3480
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3481
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3482
lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3483
  by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3484
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3485
lemma eq_fps_sin:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3486
  assumes 0: "a $ 0 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3487
    and 1: "a $ 1 = c"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3488
    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
  3489
  shows "a = fps_sin c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3490
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3491
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3492
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3493
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3494
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3495
  apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3496
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3497
  apply (subst minus_divide_left)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3498
  apply (subst eq_divide_iff)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3499
  apply (simp del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3500
  apply (simp only: mult_ac)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3501
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3502
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3503
lemma eq_fps_cos:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3504
  assumes 0: "a $ 0 = 1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3505
    and 1: "a $ 1 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3506
    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
  3507
  shows "a = fps_cos c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3508
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3509
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3510
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3511
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3512
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3513
  apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3514
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3515
  apply (subst minus_divide_left)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3516
  apply (subst eq_divide_iff)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3517
  apply (simp del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3518
  apply (simp only: mult_ac)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3519
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3520
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3521
lemma mult_nth_0 [simp]: "(a * b) $ 0 = a $ 0 * b $ 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3522
  by (simp add: fps_mult_nth)
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3523
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3524
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
  3525
  by (simp add: fps_mult_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3526
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3527
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
  3528
  apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3529
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3530
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3531
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3532
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3533
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3534
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
  3535
  apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3536
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3537
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3538
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3539
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  3540
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  3541
lemma fps_sin_even: "fps_sin (- c) = - fps_sin c"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  3542
  by (auto simp add: fps_eq_iff fps_sin_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  3543
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  3544
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
  3545
  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
  3546
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3547
definition "fps_tan c = fps_sin c / fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3548
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  3549
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
  3550
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3551
  have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3552
  show ?thesis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3553
    using fps_sin_cos_sum_of_squares[of c]
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3554
    apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3555
      fps_const_neg[symmetric] field_simps power2_eq_square del: fps_const_neg)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 49834
diff changeset
  3556
    unfolding distrib_left[symmetric]
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3557
    apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3558
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3559
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  3560
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3561
text {* Connection to E c over the complex numbers --- Euler and De Moivre*}
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3562
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
  3563
  (is "?l = ?r")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3564
proof -
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3565
  { fix n :: nat
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3566
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3567
      assume en: "even n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3568
      from en obtain m where m: "n = 2 * m"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3569
        unfolding even_mult_two_ex by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3570
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3571
      have "?l $n = ?r$n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3572
        by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3573
    }
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3574
    moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3575
    {
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3576
      assume on: "odd n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3577
      from on obtain m where m: "n = 2*m + 1"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3578
        unfolding odd_nat_equiv_def2 by (auto simp add: mult_2)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3579
      have "?l $n = ?r$n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3580
        by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3581
          power_mult power_minus)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3582
    }
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3583
    ultimately have "?l $n = ?r$n"  by blast
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3584
  } then show ?thesis by (simp add: fps_eq_iff)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3585
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3586
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3587
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
  3588
  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
  3589
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3590
lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3591
  by (simp add: fps_eq_iff fps_const_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3592
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  3593
lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a:: {comm_ring_1})"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  3594
  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
  3595
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3596
lemma fps_cos_Eii: "fps_cos c = (E (ii * c) + E (- ii * c)) / fps_const 2"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3597
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3598
  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
  3599
    by (simp add: numeral_fps_const)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3600
  show ?thesis
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3601
  unfolding Eii_sin_cos minus_mult_commute
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3602
  by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_def fps_const_inverse th)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3603
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3604
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3605
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
  3606
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3607
  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
  3608
    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
  3609
  show ?thesis
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3610
    unfolding Eii_sin_cos minus_mult_commute
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3611
    by (simp add: fps_sin_even fps_cos_odd fps_divide_def fps_const_inverse th)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3612
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3613
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3614
lemma fps_tan_Eii:
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3615
  "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
  3616
  unfolding fps_tan_def fps_sin_Eii fps_cos_Eii mult_minus_left E_neg
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3617
  apply (simp add: fps_divide_def fps_inverse_mult fps_const_mult[symmetric] fps_const_inverse del: fps_const_mult)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3618
  apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3619
  done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3620
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3621
lemma fps_demoivre: "(fps_cos a + fps_const ii * fps_sin a)^n = fps_cos (of_nat n * a) + fps_const ii * fps_sin (of_nat n * a)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3622
  unfolding Eii_sin_cos[symmetric] E_power_mult
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3623
  by (simp add: mult_ac)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3624
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3625
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3626
subsection {* Hypergeometric series *}
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3627
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3628
definition "F as bs (c::'a::{field_char_0, field_inverse_zero}) =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3629
  Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3630
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3631
lemma F_nth[simp]: "F as bs c $ n =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3632
  (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3633
    (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
  3634
  by (simp add: F_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3635
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3636
lemma foldl_mult_start:
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3637
  "foldl (%r x. r * f x) (v::'a::comm_ring_1) as * x = foldl (%r x. r * f x) (v * x) as "
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3638
  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
  3639
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3640
lemma foldr_mult_foldl: "foldr (%x r. r * f x) as v = foldl (%r x. r * f x) (v :: 'a::comm_ring_1) as"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3641
  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
  3642
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3643
lemma F_nth_alt: "F as bs c $ n = foldr (\<lambda>a r. r * pochhammer a n) as (c ^ n) /
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3644
    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
  3645
  by (simp add: foldl_mult_start foldr_mult_foldl)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3646
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3647
lemma F_E[simp]: "F [] [] c = E c"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3648
  by (simp add: fps_eq_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3649
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3650
lemma F_1_0[simp]: "F [1] [] c = 1/(1 - fps_const c * X)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3651
proof -
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3652
  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
  3653
  have th0: "(fps_const c * X) $ 0 = 0" by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3654
  show ?thesis unfolding gp[OF th0, symmetric]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3655
    by (auto simp add: fps_eq_iff pochhammer_fact[symmetric] fps_compose_nth power_mult_distrib cond_value_iff setsum_delta' cong del: if_weak_cong)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3656
qed
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3657
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3658
lemma F_B[simp]: "F [-a] [] (- 1) = fps_binomial a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3659
  by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3660
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3661
lemma F_0[simp]: "F as bs c $0 = 1"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3662
  apply simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3663
  apply (subgoal_tac "ALL as. foldl (%(r::'a) (a::'a). r) 1 as = 1")
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3664
  apply auto
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3665
  apply (induct_tac as)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3666
  apply auto
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3667
  done
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3668
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3669
lemma foldl_prod_prod: "foldl (%(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (%r x. r * g x) w as =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3670
    foldl (%r x. r * f x * g x) (v*w) as"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3671
  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
  3672
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3673
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3674
lemma F_rec: "F as bs c $ Suc n = ((foldl (%r a. r* (a + of_nat n)) c as)/ (foldl (%r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * F as bs c $ n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3675
  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
  3676
  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
  3677
  unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3678
  apply (simp add: algebra_simps of_nat_mult)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3679
  done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3680
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3681
lemma XD_nth[simp]: "XD a $ n = (if n=0 then 0 else of_nat n * a$n)"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3682
  by (simp add: XD_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3683
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3684
lemma XD_0th[simp]: "XD a $ 0 = 0" by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3685
lemma XD_Suc[simp]:" XD a $ Suc n = of_nat (Suc n) * a $ Suc n" by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3686
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3687
definition "XDp c a = XD a + fps_const c * a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3688
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3689
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
  3690
  by (simp add: XDp_def algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3691
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3692
lemma XDp_commute: "XDp b o XDp (c::'a::comm_ring_1) = XDp c o 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
  3693
  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
  3694
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3695
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
  3696
  by (simp add: fun_eq_iff fps_eq_iff)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3697
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3698
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
  3699
  by (simp add: fps_eq_iff fps_integral_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3700
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3701
lemma F_minus_nat:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3702
  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, field_inverse_zero}) $ k =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3703
    (if k <= n then
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3704
      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
  3705
     else 0)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3706
  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, field_inverse_zero}) $ k =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3707
    (if k <= m then
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3708
      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
  3709
     else 0)"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3710
  by (auto simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3711
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3712
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
  3713
  apply simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3714
  apply (subst setsum_insert[symmetric])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3715
  apply (auto simp add: not_less setsum_head_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3716
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3717
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3718
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
  3719
  by (cases n) (simp_all add: pochhammer_rec)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3720
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3721
lemma XDp_foldr_nth [simp]: "foldr (%c r. XDp c o r) cs (%c. XDp c a) c0 $ n =
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3722
  foldr (%c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3723
  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
  3724
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3725
lemma genric_XDp_foldr_nth:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3726
  assumes f: "ALL n c a. f c a $ n = (of_nat n + k c) * a$n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3727
  shows "foldr (%c r. f c o r) cs (%c. g c a) c0 $ n =
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3728
    foldr (%c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3729
  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
  3730
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3731
lemma dist_less_imp_nth_equal:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3732
  assumes "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3733
    and"j \<le> i"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3734
  shows "f $ j = g $ j"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3735
proof (cases "f = g")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3736
  case False
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3737
  hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3738
  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3739
    by (simp add: split_if_asm dist_fps_def)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3740
  moreover
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3741
  from fps_eq_least_unique[OF `f \<noteq> g`]
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3742
  obtain n where "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3743
  moreover hence "\<And>m. m < n \<Longrightarrow> f$m = g$m" "f$n \<noteq> g$n"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3744
    by (auto simp add: leastP_def setge_def)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3745
  ultimately show ?thesis using `j \<le> i` by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3746
next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3747
  case True
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3748
  then show ?thesis by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3749
qed
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3750
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3751
lemma nth_equal_imp_dist_less:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3752
  assumes "\<And>j. j \<le> i \<Longrightarrow> f $ j = g $ j"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3753
  shows "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3754
proof (cases "f = g")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3755
  case False
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3756
  hence "\<exists>n. f $ n \<noteq> g $ n" by (simp add: fps_eq_iff)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3757
  with assms have "dist f g = inverse (2 ^ (The (leastP (\<lambda>n. f $ n \<noteq> g $ n))))"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3758
    by (simp add: split_if_asm dist_fps_def)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3759
  moreover
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3760
  from fps_eq_least_unique[OF `f \<noteq> g`]
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3761
  obtain n where "leastP (\<lambda>n. f$n \<noteq> g$n) n" "The (leastP (\<lambda>n. f $ n \<noteq> g $ n)) = n" by blast
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3762
  with assms have "i < The (leastP (\<lambda>n. f $ n \<noteq> g $ n))"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3763
    by (metis (full_types) leastPD1 not_le)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3764
  ultimately show ?thesis by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3765
next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3766
  case True
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3767
  then show ?thesis by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3768
qed
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3769
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3770
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
  3771
  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
  3772
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3773
instance fps :: (comm_ring_1) complete_space
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3774
proof
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3775
  fix X::"nat \<Rightarrow> 'a fps"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3776
  assume "Cauchy X"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3777
  {
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3778
    fix i
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3779
    have "0 < inverse ((2::real)^i)" by simp
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3780
    from metric_CauchyD[OF `Cauchy X` this] dist_less_imp_nth_equal
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3781
    have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. X M $ j = X m $ j" by blast
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3782
  }
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3783
  then obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j" by metis
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3784
  hence "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j" by metis
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3785
  show "convergent X"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3786
  proof (rule convergentI)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3787
    show "X ----> Abs_fps (\<lambda>i. X (M i) $ i)"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3788
      unfolding tendsto_iff
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3789
    proof safe
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3790
      fix e::real assume "0 < e"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3791
      with LIMSEQ_inverse_realpow_zero[of 2, simplified, simplified filterlim_iff,
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3792
        THEN spec, of "\<lambda>x. x < e"]
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3793
      have "eventually (\<lambda>i. inverse (2 ^ i) < e) sequentially"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3794
        apply safe
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3795
        apply (auto simp: eventually_nhds)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3796
        done
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3797
      then obtain i where "inverse (2 ^ i) < e" by (auto simp: eventually_sequentially)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3798
      have "eventually (\<lambda>x. M i \<le> x) sequentially" by (auto simp: eventually_sequentially)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3799
      thus "eventually (\<lambda>x. dist (X x) (Abs_fps (\<lambda>i. X (M i) $ i)) < e) sequentially"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3800
      proof eventually_elim
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3801
        fix x
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3802
        assume "M i \<le> x"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3803
        moreover
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3804
        have "\<And>j. j \<le> i \<Longrightarrow> X (M i) $ j = X (M j) $ j"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3805
          using M by (metis nat_le_linear)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3806
        ultimately have "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < inverse (2 ^ i)"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3807
          using M by (force simp: dist_less_eq_nth_equal)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3808
        also note `inverse (2 ^ i) < e`
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3809
        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
  3810
      qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3811
    qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3812
  qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3813
qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  3814
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  3815
end