| author | wenzelm | 
| Sun, 27 Dec 2020 15:15:37 +0100 | |
| changeset 73013 | d4b67dc6f4eb | 
| parent 70045 | 7b6add80e3a5 | 
| child 73655 | 26a1d66b9077 | 
| permissions | -rw-r--r-- | 
| 70043 | 1  | 
(* Author: Andreas Lochbihler, ETH Zurich  | 
2  | 
Author: Florian Haftmann, TU Muenchen  | 
|
3  | 
with some material ported from HOL Light by LCP  | 
|
4  | 
*)  | 
|
5  | 
||
6  | 
section \<open>Polynomial mapping: combination of almost everywhere zero functions with an algebraic view\<close>  | 
|
7  | 
||
8  | 
theory Poly_Mapping  | 
|
9  | 
imports Groups_Big_Fun Fun_Lexorder More_List  | 
|
10  | 
begin  | 
|
11  | 
||
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
12  | 
subsection \<open>Preliminary: auxiliary operations for \emph{almost everywhere zero}\<close>
 | 
| 70043 | 13  | 
|
14  | 
text \<open>  | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
15  | 
  A central notion for polynomials are functions being \emph{almost everywhere zero}.
 | 
| 70043 | 16  | 
For these we provide some auxiliary definitions and lemmas.  | 
17  | 
\<close>  | 
|
18  | 
||
19  | 
lemma finite_mult_not_eq_zero_leftI:  | 
|
20  | 
fixes f :: "'b \<Rightarrow> 'a :: mult_zero"  | 
|
21  | 
  assumes "finite {a. f a \<noteq> 0}"
 | 
|
22  | 
  shows "finite {a. g a * f a \<noteq> 0}"
 | 
|
23  | 
proof -  | 
|
24  | 
  have "{a. g a * f a \<noteq> 0} \<subseteq> {a. f a \<noteq> 0}" by auto
 | 
|
25  | 
then show ?thesis using assms by (rule finite_subset)  | 
|
26  | 
qed  | 
|
27  | 
||
28  | 
lemma finite_mult_not_eq_zero_rightI:  | 
|
29  | 
fixes f :: "'b \<Rightarrow> 'a :: mult_zero"  | 
|
30  | 
  assumes "finite {a. f a \<noteq> 0}"
 | 
|
31  | 
  shows "finite {a. f a * g a \<noteq> 0}"
 | 
|
32  | 
proof -  | 
|
33  | 
  have "{a. f a * g a \<noteq> 0} \<subseteq> {a. f a \<noteq> 0}" by auto
 | 
|
34  | 
then show ?thesis using assms by (rule finite_subset)  | 
|
35  | 
qed  | 
|
36  | 
||
37  | 
lemma finite_mult_not_eq_zero_prodI:  | 
|
38  | 
fixes f g :: "'a \<Rightarrow> 'b::semiring_0"  | 
|
39  | 
  assumes "finite {a. f a \<noteq> 0}" (is "finite ?F")
 | 
|
40  | 
  assumes "finite {b. g b \<noteq> 0}" (is "finite ?G")
 | 
|
41  | 
  shows "finite {(a, b). f a * g b \<noteq> 0}"
 | 
|
42  | 
proof -  | 
|
43  | 
from assms have "finite (?F \<times> ?G)"  | 
|
44  | 
by blast  | 
|
45  | 
  then have "finite {(a, b). f a \<noteq> 0 \<and> g b \<noteq> 0}"
 | 
|
46  | 
by simp  | 
|
47  | 
then show ?thesis  | 
|
48  | 
by (rule rev_finite_subset) auto  | 
|
49  | 
qed  | 
|
50  | 
||
51  | 
lemma finite_not_eq_zero_sumI:  | 
|
52  | 
fixes f g :: "'a::monoid_add \<Rightarrow> 'b::semiring_0"  | 
|
53  | 
  assumes "finite {a. f a \<noteq> 0}" (is "finite ?F")
 | 
|
54  | 
  assumes "finite {b. g b \<noteq> 0}" (is "finite ?G")
 | 
|
55  | 
  shows "finite {a + b | a b. f a \<noteq> 0 \<and> g b \<noteq> 0}" (is "finite ?FG")
 | 
|
56  | 
proof -  | 
|
57  | 
from assms have "finite (?F \<times> ?G)"  | 
|
58  | 
by (simp add: finite_cartesian_product_iff)  | 
|
59  | 
then have "finite (case_prod plus ` (?F \<times> ?G))"  | 
|
60  | 
by (rule finite_imageI)  | 
|
61  | 
also have "case_prod plus ` (?F \<times> ?G) = ?FG"  | 
|
62  | 
by auto  | 
|
63  | 
finally show ?thesis  | 
|
64  | 
by simp  | 
|
65  | 
qed  | 
|
66  | 
||
67  | 
lemma finite_mult_not_eq_zero_sumI:  | 
|
68  | 
fixes f g :: "'a::monoid_add \<Rightarrow> 'b::semiring_0"  | 
|
69  | 
  assumes "finite {a. f a \<noteq> 0}"
 | 
|
70  | 
  assumes "finite {b. g b \<noteq> 0}"
 | 
|
71  | 
  shows "finite {a + b | a b. f a * g b \<noteq> 0}"
 | 
|
72  | 
proof -  | 
|
73  | 
from assms  | 
|
74  | 
  have "finite {a + b | a b. f a \<noteq> 0 \<and> g b \<noteq> 0}"
 | 
|
75  | 
by (rule finite_not_eq_zero_sumI)  | 
|
76  | 
then show ?thesis  | 
|
77  | 
by (rule rev_finite_subset) (auto dest: mult_not_zero)  | 
|
78  | 
qed  | 
|
79  | 
||
80  | 
lemma finite_Sum_any_not_eq_zero_weakenI:  | 
|
81  | 
  assumes "finite {a. \<exists>b. f a b \<noteq> 0}"
 | 
|
82  | 
  shows "finite {a. Sum_any (f a) \<noteq> 0}"
 | 
|
83  | 
proof -  | 
|
84  | 
  have "{a. Sum_any (f a) \<noteq> 0} \<subseteq> {a. \<exists>b. f a b \<noteq> 0}"
 | 
|
85  | 
by (auto elim: Sum_any.not_neutral_obtains_not_neutral)  | 
|
86  | 
then show ?thesis using assms by (rule finite_subset)  | 
|
87  | 
qed  | 
|
88  | 
||
89  | 
context zero  | 
|
90  | 
begin  | 
|
91  | 
||
92  | 
definition "when" :: "'a \<Rightarrow> bool \<Rightarrow> 'a" (infixl "when" 20)  | 
|
93  | 
where  | 
|
94  | 
"(a when P) = (if P then a else 0)"  | 
|
95  | 
||
96  | 
text \<open>  | 
|
97  | 
Case distinctions always complicate matters, particularly when  | 
|
98  | 
  nested.  The @{const "when"} operation allows to minimise these
 | 
|
99  | 
  if @{term 0} is the false-case value and makes proof obligations
 | 
|
100  | 
much more readable.  | 
|
101  | 
\<close>  | 
|
102  | 
||
103  | 
lemma "when" [simp]:  | 
|
104  | 
"P \<Longrightarrow> (a when P) = a"  | 
|
105  | 
"\<not> P \<Longrightarrow> (a when P) = 0"  | 
|
106  | 
by (simp_all add: when_def)  | 
|
107  | 
||
108  | 
lemma when_simps [simp]:  | 
|
109  | 
"(a when True) = a"  | 
|
110  | 
"(a when False) = 0"  | 
|
111  | 
by simp_all  | 
|
112  | 
||
113  | 
lemma when_cong:  | 
|
114  | 
assumes "P \<longleftrightarrow> Q"  | 
|
115  | 
and "Q \<Longrightarrow> a = b"  | 
|
116  | 
shows "(a when P) = (b when Q)"  | 
|
117  | 
using assms by (simp add: when_def)  | 
|
118  | 
||
119  | 
lemma zero_when [simp]:  | 
|
120  | 
"(0 when P) = 0"  | 
|
121  | 
by (simp add: when_def)  | 
|
122  | 
||
123  | 
lemma when_when:  | 
|
124  | 
"(a when P when Q) = (a when P \<and> Q)"  | 
|
125  | 
by (cases Q) simp_all  | 
|
126  | 
||
127  | 
lemma when_commute:  | 
|
128  | 
"(a when Q when P) = (a when P when Q)"  | 
|
129  | 
by (simp add: when_when conj_commute)  | 
|
130  | 
||
131  | 
lemma when_neq_zero [simp]:  | 
|
132  | 
"(a when P) \<noteq> 0 \<longleftrightarrow> P \<and> a \<noteq> 0"  | 
|
133  | 
by (cases P) simp_all  | 
|
134  | 
||
135  | 
end  | 
|
136  | 
||
137  | 
context monoid_add  | 
|
138  | 
begin  | 
|
139  | 
||
140  | 
lemma when_add_distrib:  | 
|
141  | 
"(a + b when P) = (a when P) + (b when P)"  | 
|
142  | 
by (simp add: when_def)  | 
|
143  | 
||
144  | 
end  | 
|
145  | 
||
146  | 
context semiring_1  | 
|
147  | 
begin  | 
|
148  | 
||
149  | 
lemma zero_power_eq:  | 
|
150  | 
"0 ^ n = (1 when n = 0)"  | 
|
151  | 
by (simp add: power_0_left)  | 
|
152  | 
||
153  | 
end  | 
|
154  | 
||
155  | 
context comm_monoid_add  | 
|
156  | 
begin  | 
|
157  | 
||
158  | 
lemma Sum_any_when_equal [simp]:  | 
|
159  | 
"(\<Sum>a. (f a when a = b)) = f b"  | 
|
160  | 
by (simp add: when_def)  | 
|
161  | 
||
162  | 
lemma Sum_any_when_equal' [simp]:  | 
|
163  | 
"(\<Sum>a. (f a when b = a)) = f b"  | 
|
164  | 
by (simp add: when_def)  | 
|
165  | 
||
166  | 
lemma Sum_any_when_independent:  | 
|
167  | 
"(\<Sum>a. g a when P) = ((\<Sum>a. g a) when P)"  | 
|
168  | 
by (cases P) simp_all  | 
|
169  | 
||
170  | 
lemma Sum_any_when_dependent_prod_right:  | 
|
171  | 
"(\<Sum>(a, b). g a when b = h a) = (\<Sum>a. g a)"  | 
|
172  | 
proof -  | 
|
173  | 
  have "inj_on (\<lambda>a. (a, h a)) {a. g a \<noteq> 0}"
 | 
|
174  | 
by (rule inj_onI) auto  | 
|
175  | 
then show ?thesis unfolding Sum_any.expand_set  | 
|
176  | 
by (rule sum.reindex_cong) auto  | 
|
177  | 
qed  | 
|
178  | 
||
179  | 
lemma Sum_any_when_dependent_prod_left:  | 
|
180  | 
"(\<Sum>(a, b). g b when a = h b) = (\<Sum>b. g b)"  | 
|
181  | 
proof -  | 
|
182  | 
have "(\<Sum>(a, b). g b when a = h b) = (\<Sum>(b, a). g b when a = h b)"  | 
|
183  | 
by (rule Sum_any.reindex_cong [of prod.swap]) (simp_all add: fun_eq_iff)  | 
|
184  | 
then show ?thesis by (simp add: Sum_any_when_dependent_prod_right)  | 
|
185  | 
qed  | 
|
186  | 
||
187  | 
end  | 
|
188  | 
||
189  | 
context cancel_comm_monoid_add  | 
|
190  | 
begin  | 
|
191  | 
||
192  | 
lemma when_diff_distrib:  | 
|
193  | 
"(a - b when P) = (a when P) - (b when P)"  | 
|
194  | 
by (simp add: when_def)  | 
|
195  | 
||
196  | 
end  | 
|
197  | 
||
198  | 
context group_add  | 
|
199  | 
begin  | 
|
200  | 
||
201  | 
lemma when_uminus_distrib:  | 
|
202  | 
"(- a when P) = - (a when P)"  | 
|
203  | 
by (simp add: when_def)  | 
|
204  | 
||
205  | 
end  | 
|
206  | 
||
207  | 
context mult_zero  | 
|
208  | 
begin  | 
|
209  | 
||
210  | 
lemma mult_when:  | 
|
211  | 
"a * (b when P) = (a * b when P)"  | 
|
212  | 
by (cases P) simp_all  | 
|
213  | 
||
214  | 
lemma when_mult:  | 
|
215  | 
"(a when P) * b = (a * b when P)"  | 
|
216  | 
by (cases P) simp_all  | 
|
217  | 
||
218  | 
end  | 
|
219  | 
||
220  | 
||
221  | 
subsection \<open>Type definition\<close>  | 
|
222  | 
||
223  | 
text \<open>  | 
|
224  | 
The following type is of central importance:  | 
|
225  | 
\<close>  | 
|
226  | 
||
227  | 
typedef (overloaded) ('a, 'b) poly_mapping ("(_ \<Rightarrow>\<^sub>0 /_)" [1, 0] 0) =
 | 
|
228  | 
  "{f :: 'a \<Rightarrow> 'b::zero. finite {x. f x \<noteq> 0}}"
 | 
|
229  | 
morphisms lookup Abs_poly_mapping  | 
|
230  | 
proof -  | 
|
231  | 
have "(\<lambda>_::'a. (0 :: 'b)) \<in> ?poly_mapping" by simp  | 
|
232  | 
then show ?thesis by (blast intro!: exI)  | 
|
233  | 
qed  | 
|
234  | 
||
235  | 
declare lookup_inverse [simp]  | 
|
236  | 
declare lookup_inject [simp]  | 
|
237  | 
||
238  | 
lemma lookup_Abs_poly_mapping [simp]:  | 
|
239  | 
  "finite {x. f x \<noteq> 0} \<Longrightarrow> lookup (Abs_poly_mapping f) = f"
 | 
|
240  | 
using Abs_poly_mapping_inverse [of f] by simp  | 
|
241  | 
||
242  | 
lemma finite_lookup [simp]:  | 
|
243  | 
  "finite {k. lookup f k \<noteq> 0}"
 | 
|
244  | 
using poly_mapping.lookup [of f] by simp  | 
|
245  | 
||
246  | 
lemma finite_lookup_nat [simp]:  | 
|
247  | 
fixes f :: "'a \<Rightarrow>\<^sub>0 nat"  | 
|
248  | 
  shows "finite {k. 0 < lookup f k}"
 | 
|
249  | 
using poly_mapping.lookup [of f] by simp  | 
|
250  | 
||
251  | 
lemma poly_mapping_eqI:  | 
|
252  | 
assumes "\<And>k. lookup f k = lookup g k"  | 
|
253  | 
shows "f = g"  | 
|
254  | 
using assms unfolding poly_mapping.lookup_inject [symmetric]  | 
|
255  | 
by blast  | 
|
256  | 
||
257  | 
lemma poly_mapping_eq_iff: "a = b \<longleftrightarrow> lookup a = lookup b"  | 
|
258  | 
by auto  | 
|
259  | 
||
260  | 
text \<open>  | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
261  | 
  We model the universe of functions being \emph{almost everywhere zero}
 | 
| 70043 | 262  | 
  by means of a separate type @{typ "('a, 'b) poly_mapping"}.
 | 
263  | 
For convenience we provide a suggestive infix syntax which  | 
|
264  | 
is a variant of the usual function space syntax. Conversion between both types  | 
|
265  | 
happens through the morphisms  | 
|
266  | 
  \begin{quote}
 | 
|
267  | 
    @{term_type lookup}
 | 
|
268  | 
  \end{quote}
 | 
|
269  | 
  \begin{quote}
 | 
|
270  | 
    @{term_type Abs_poly_mapping}
 | 
|
271  | 
  \end{quote}
 | 
|
272  | 
satisfying  | 
|
273  | 
  \begin{quote}
 | 
|
274  | 
    @{thm lookup_inverse}
 | 
|
275  | 
  \end{quote}
 | 
|
276  | 
  \begin{quote}
 | 
|
277  | 
    @{thm lookup_Abs_poly_mapping}
 | 
|
278  | 
  \end{quote}
 | 
|
279  | 
Luckily, we have rarely to deal with those low-level morphisms explicitly  | 
|
280  | 
  but rely on Isabelle's \emph{lifting} package with its method \<open>transfer\<close>
 | 
|
281  | 
and its specification tool \<open>lift_definition\<close>.  | 
|
282  | 
\<close>  | 
|
283  | 
||
284  | 
setup_lifting type_definition_poly_mapping  | 
|
285  | 
code_datatype Abs_poly_mapping\<comment>\<open>FIXME? workaround for preventing \<open>code_abstype\<close> setup\<close>  | 
|
286  | 
||
287  | 
text \<open>  | 
|
288  | 
  @{typ "'a \<Rightarrow>\<^sub>0 'b"} serves distinctive purposes:
 | 
|
289  | 
  \begin{enumerate}
 | 
|
290  | 
    \item A clever nesting as @{typ "(nat \<Rightarrow>\<^sub>0 nat) \<Rightarrow>\<^sub>0 'a"}
 | 
|
291  | 
later in theory \<open>MPoly\<close> gives a suitable  | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
292  | 
      representation type for polynomials \emph{almost for free}:
 | 
| 70043 | 293  | 
      Interpreting @{typ "nat \<Rightarrow>\<^sub>0 nat"} as a mapping from variable identifiers
 | 
294  | 
to exponents yields monomials, and the whole type maps monomials  | 
|
295  | 
      to coefficients.  Lets call this the \emph{ultimate interpretation}.
 | 
|
296  | 
    \item A further more specialised type isomorphic to @{typ "nat \<Rightarrow>\<^sub>0 'a"}
 | 
|
297  | 
is apt to direct implementation using code generation  | 
|
298  | 
      \cite{Haftmann-Nipkow:2010:code}.
 | 
|
299  | 
  \end{enumerate}
 | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
300  | 
  Note that despite the names \emph{mapping} and \emph{lookup} suggest something
 | 
| 70043 | 301  | 
  implementation-near, it is best to keep @{typ "'a \<Rightarrow>\<^sub>0 'b"} as an abstract
 | 
302  | 
  \emph{algebraic} type
 | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
303  | 
  providing operations like \emph{addition}, \emph{multiplication} without any notion
 | 
| 70043 | 304  | 
of key-order, data structures etc. This implementations-specific notions are  | 
305  | 
easily introduced later for particular implementations but do not provide any  | 
|
306  | 
gain for specifying logical properties of polynomials.  | 
|
307  | 
\<close>  | 
|
308  | 
||
309  | 
subsection \<open>Additive structure\<close>  | 
|
310  | 
||
311  | 
text \<open>  | 
|
312  | 
The additive structure covers the usual operations \<open>0\<close>, \<open>+\<close> and  | 
|
313  | 
(unary and binary) \<open>-\<close>. Recalling the ultimate interpretation, it  | 
|
314  | 
is obvious that these have just lift the corresponding operations on values  | 
|
315  | 
to mappings.  | 
|
316  | 
||
317  | 
Isabelle has a rich hierarchy of algebraic type classes, and in such situations  | 
|
318  | 
of pointwise lifting a typical pattern is to have instantiations for a considerable  | 
|
319  | 
number of type classes.  | 
|
320  | 
||
321  | 
The operations themselves are specified using \<open>lift_definition\<close>, where  | 
|
| 
70045
 
7b6add80e3a5
fixed markup in Poly_Mapping; Free_Abelian_Groups (but not yet imported by Algebra!)
 
paulson <lp15@cam.ac.uk> 
parents: 
70043 
diff
changeset
 | 
322  | 
  the proofs of the \emph{almost everywhere zero} property can be significantly involved.
 | 
| 70043 | 323  | 
|
324  | 
  The @{const lookup} operation is supposed to be usable explicitly (unless in
 | 
|
325  | 
other situations where the morphisms between types are somehow internal  | 
|
326  | 
  to the \emph{lifting} package).  Hence it is good style to provide explicit rewrite
 | 
|
327  | 
  rules how @{const lookup} acts on operations immediately.
 | 
|
328  | 
\<close>  | 
|
329  | 
||
330  | 
instantiation poly_mapping :: (type, zero) zero  | 
|
331  | 
begin  | 
|
332  | 
||
333  | 
lift_definition zero_poly_mapping :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
334  | 
is "\<lambda>k. 0"  | 
|
335  | 
by simp  | 
|
336  | 
||
337  | 
instance ..  | 
|
338  | 
||
339  | 
end  | 
|
340  | 
||
341  | 
lemma Abs_poly_mapping [simp]: "Abs_poly_mapping (\<lambda>k. 0) = 0"  | 
|
342  | 
by (simp add: zero_poly_mapping.abs_eq)  | 
|
343  | 
||
344  | 
lemma lookup_zero [simp]: "lookup 0 k = 0"  | 
|
345  | 
by transfer rule  | 
|
346  | 
||
347  | 
instantiation poly_mapping :: (type, monoid_add) monoid_add  | 
|
348  | 
begin  | 
|
349  | 
||
350  | 
lift_definition plus_poly_mapping ::  | 
|
351  | 
    "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
 | 
|
352  | 
is "\<lambda>f1 f2 k. f1 k + f2 k"  | 
|
353  | 
proof -  | 
|
354  | 
fix f1 f2 :: "'a \<Rightarrow> 'b"  | 
|
355  | 
  assume "finite {k. f1 k \<noteq> 0}"
 | 
|
356  | 
    and "finite {k. f2 k \<noteq> 0}"
 | 
|
357  | 
  then have "finite ({k. f1 k \<noteq> 0} \<union> {k. f2 k \<noteq> 0})" by auto
 | 
|
358  | 
  moreover have "{x. f1 x + f2 x \<noteq> 0} \<subseteq> {k. f1 k \<noteq> 0} \<union> {k. f2 k \<noteq> 0}"
 | 
|
359  | 
by auto  | 
|
360  | 
  ultimately show "finite {x. f1 x + f2 x \<noteq> 0}"
 | 
|
361  | 
by (blast intro: finite_subset)  | 
|
362  | 
qed  | 
|
363  | 
||
364  | 
instance  | 
|
365  | 
by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+  | 
|
366  | 
||
367  | 
end  | 
|
368  | 
||
369  | 
lemma lookup_add:  | 
|
370  | 
"lookup (f + g) k = lookup f k + lookup g k"  | 
|
371  | 
by transfer rule  | 
|
372  | 
||
373  | 
instance poly_mapping :: (type, comm_monoid_add) comm_monoid_add  | 
|
374  | 
by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+  | 
|
375  | 
||
376  | 
lemma lookup_sum: "lookup (sum pp X) i = sum (\<lambda>x. lookup (pp x) i) X"  | 
|
377  | 
by (induction rule: infinite_finite_induct) (auto simp: lookup_add)  | 
|
378  | 
||
379  | 
(*instance poly_mapping :: (type, "{monoid_add, cancel_semigroup_add}") cancel_semigroup_add
 | 
|
380  | 
by intro_classes (transfer, simp add: fun_eq_iff)+*)  | 
|
381  | 
||
382  | 
instantiation poly_mapping :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add  | 
|
383  | 
begin  | 
|
384  | 
||
385  | 
lift_definition minus_poly_mapping :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
 | 
|
386  | 
is "\<lambda>f1 f2 k. f1 k - f2 k"  | 
|
387  | 
proof -  | 
|
388  | 
fix f1 f2 :: "'a \<Rightarrow> 'b"  | 
|
389  | 
  assume "finite {k. f1 k \<noteq> 0}"
 | 
|
390  | 
    and "finite {k. f2 k \<noteq> 0}"
 | 
|
391  | 
  then have "finite ({k. f1 k \<noteq> 0} \<union> {k. f2 k \<noteq> 0})" by auto
 | 
|
392  | 
  moreover have "{x. f1 x - f2 x \<noteq> 0} \<subseteq> {k. f1 k \<noteq> 0} \<union> {k. f2 k \<noteq> 0}"
 | 
|
393  | 
by auto  | 
|
394  | 
  ultimately show "finite {x. f1 x - f2 x \<noteq> 0}" by (blast intro: finite_subset)
 | 
|
395  | 
qed  | 
|
396  | 
||
397  | 
instance  | 
|
398  | 
by intro_classes (transfer, simp add: fun_eq_iff diff_diff_add)+  | 
|
399  | 
||
400  | 
end  | 
|
401  | 
||
402  | 
instantiation poly_mapping :: (type, ab_group_add) ab_group_add  | 
|
403  | 
begin  | 
|
404  | 
||
405  | 
lift_definition uminus_poly_mapping :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
 | 
|
406  | 
is uminus  | 
|
407  | 
by simp  | 
|
408  | 
||
409  | 
||
410  | 
instance  | 
|
411  | 
by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+  | 
|
412  | 
||
413  | 
end  | 
|
414  | 
||
415  | 
lemma lookup_uminus [simp]:  | 
|
416  | 
"lookup (- f) k = - lookup f k"  | 
|
417  | 
by transfer simp  | 
|
418  | 
||
419  | 
lemma lookup_minus:  | 
|
420  | 
"lookup (f - g) k = lookup f k - lookup g k"  | 
|
421  | 
by transfer rule  | 
|
422  | 
||
423  | 
||
424  | 
subsection \<open>Multiplicative structure\<close>  | 
|
425  | 
||
426  | 
instantiation poly_mapping :: (zero, zero_neq_one) zero_neq_one  | 
|
427  | 
begin  | 
|
428  | 
||
429  | 
lift_definition one_poly_mapping :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
430  | 
is "\<lambda>k. 1 when k = 0"  | 
|
431  | 
by simp  | 
|
432  | 
||
433  | 
instance  | 
|
434  | 
by intro_classes (transfer, simp add: fun_eq_iff)  | 
|
435  | 
||
436  | 
end  | 
|
437  | 
||
438  | 
lemma lookup_one:  | 
|
439  | 
"lookup 1 k = (1 when k = 0)"  | 
|
440  | 
by transfer rule  | 
|
441  | 
||
442  | 
lemma lookup_one_zero [simp]:  | 
|
443  | 
"lookup 1 0 = 1"  | 
|
444  | 
by transfer simp  | 
|
445  | 
||
446  | 
definition prod_fun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a::monoid_add \<Rightarrow> 'b::semiring_0"
 | 
|
447  | 
where  | 
|
448  | 
"prod_fun f1 f2 k = (\<Sum>l. f1 l * (\<Sum>q. (f2 q when k = l + q)))"  | 
|
449  | 
||
450  | 
lemma prod_fun_unfold_prod:  | 
|
451  | 
fixes f g :: "'a :: monoid_add \<Rightarrow> 'b::semiring_0"  | 
|
452  | 
  assumes fin_f: "finite {a. f a \<noteq> 0}"
 | 
|
453  | 
  assumes fin_g: "finite {b. g b \<noteq> 0}"
 | 
|
454  | 
shows "prod_fun f g k = (\<Sum>(a, b). f a * g b when k = a + b)"  | 
|
455  | 
proof -  | 
|
456  | 
  let ?C = "{a. f a \<noteq> 0} \<times> {b. g b \<noteq> 0}"
 | 
|
457  | 
from fin_f fin_g have "finite ?C" by blast  | 
|
458  | 
  moreover have "{a. \<exists>b. (f a * g b when k = a + b) \<noteq> 0} \<times>
 | 
|
459  | 
    {b. \<exists>a. (f a * g b when k = a + b) \<noteq> 0} \<subseteq> {a. f a \<noteq> 0} \<times> {b. g b \<noteq> 0}"
 | 
|
460  | 
by auto  | 
|
461  | 
ultimately show ?thesis using fin_g  | 
|
462  | 
by (auto simp add: prod_fun_def  | 
|
463  | 
      Sum_any.cartesian_product [of "{a. f a \<noteq> 0} \<times> {b. g b \<noteq> 0}"] Sum_any_right_distrib mult_when)
 | 
|
464  | 
qed  | 
|
465  | 
||
466  | 
lemma finite_prod_fun:  | 
|
467  | 
fixes f1 f2 :: "'a :: monoid_add \<Rightarrow> 'b :: semiring_0"  | 
|
468  | 
  assumes fin1: "finite {l. f1 l \<noteq> 0}"
 | 
|
469  | 
  and fin2: "finite {q. f2 q \<noteq> 0}"
 | 
|
470  | 
  shows "finite {k. prod_fun f1 f2 k \<noteq> 0}"
 | 
|
471  | 
proof -  | 
|
472  | 
  have *: "finite {k. (\<exists>l. f1 l \<noteq> 0 \<and> (\<exists>q. f2 q \<noteq> 0 \<and> k = l + q))}"
 | 
|
473  | 
using assms by simp  | 
|
474  | 
  { fix k l
 | 
|
475  | 
    have "{q. (f2 q when k = l + q) \<noteq> 0} \<subseteq> {q. f2 q \<noteq> 0 \<and> k = l + q}" by auto
 | 
|
476  | 
    with fin2 have "sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} = (\<Sum>q. (f2 q when k = l + q))"
 | 
|
477  | 
      by (simp add: Sum_any.expand_superset [of "{q. f2 q \<noteq> 0 \<and> k = l + q}"]) }
 | 
|
478  | 
note aux = this  | 
|
479  | 
  have "{k. (\<Sum>l. f1 l * sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q}) \<noteq> 0}
 | 
|
480  | 
    \<subseteq> {k. (\<exists>l. f1 l * sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} \<noteq> 0)}"
 | 
|
481  | 
by (auto elim!: Sum_any.not_neutral_obtains_not_neutral)  | 
|
482  | 
  also have "\<dots> \<subseteq> {k. (\<exists>l. f1 l \<noteq> 0 \<and> sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} \<noteq> 0)}"
 | 
|
483  | 
by (auto dest: mult_not_zero)  | 
|
484  | 
  also have "\<dots> \<subseteq> {k. (\<exists>l. f1 l \<noteq> 0 \<and> (\<exists>q. f2 q \<noteq> 0 \<and> k = l + q))}"
 | 
|
485  | 
by (auto elim!: sum.not_neutral_contains_not_neutral)  | 
|
486  | 
  finally have "finite {k. (\<Sum>l. f1 l * sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q}) \<noteq> 0}"
 | 
|
487  | 
using * by (rule finite_subset)  | 
|
488  | 
  with aux have "finite {k. (\<Sum>l. f1 l * (\<Sum>q. (f2 q when k = l + q))) \<noteq> 0}"
 | 
|
489  | 
by simp  | 
|
490  | 
with fin2 show ?thesis  | 
|
491  | 
by (simp add: prod_fun_def)  | 
|
492  | 
qed  | 
|
493  | 
||
494  | 
instantiation poly_mapping :: (monoid_add, semiring_0) semiring_0  | 
|
495  | 
begin  | 
|
496  | 
||
497  | 
lift_definition times_poly_mapping :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
 | 
|
498  | 
is prod_fun  | 
|
499  | 
by(rule finite_prod_fun)  | 
|
500  | 
||
501  | 
instance  | 
|
502  | 
proof  | 
|
503  | 
fix a b c :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
504  | 
show "a * b * c = a * (b * c)"  | 
|
505  | 
proof transfer  | 
|
506  | 
fix f g h :: "'a \<Rightarrow> 'b"  | 
|
507  | 
    assume fin_f: "finite {a. f a \<noteq> 0}" (is "finite ?F")
 | 
|
508  | 
    assume fin_g: "finite {b. g b \<noteq> 0}" (is "finite ?G")
 | 
|
509  | 
    assume fin_h: "finite {c. h c \<noteq> 0}" (is "finite ?H")
 | 
|
510  | 
    from fin_f fin_g have fin_fg: "finite {(a, b). f a * g b \<noteq> 0}" (is "finite ?FG")
 | 
|
511  | 
by (rule finite_mult_not_eq_zero_prodI)  | 
|
512  | 
    from fin_g fin_h have fin_gh: "finite {(b, c). g b * h c \<noteq> 0}" (is "finite ?GH")
 | 
|
513  | 
by (rule finite_mult_not_eq_zero_prodI)  | 
|
514  | 
    from fin_f fin_g have fin_fg': "finite {a + b | a b. f a * g b \<noteq> 0}" (is "finite ?FG'")
 | 
|
515  | 
by (rule finite_mult_not_eq_zero_sumI)  | 
|
516  | 
    then have fin_fg'': "finite {d. (\<Sum>(a, b). f a * g b when d = a + b) \<noteq> 0}"
 | 
|
517  | 
by (auto intro: finite_Sum_any_not_eq_zero_weakenI)  | 
|
518  | 
    from fin_g fin_h have fin_gh': "finite {b + c | b c. g b * h c \<noteq> 0}" (is "finite ?GH'")
 | 
|
519  | 
by (rule finite_mult_not_eq_zero_sumI)  | 
|
520  | 
    then have fin_gh'': "finite {d. (\<Sum>(b, c). g b * h c when d = b + c) \<noteq> 0}"
 | 
|
521  | 
by (auto intro: finite_Sum_any_not_eq_zero_weakenI)  | 
|
522  | 
show "prod_fun (prod_fun f g) h = prod_fun f (prod_fun g h)" (is "?lhs = ?rhs")  | 
|
523  | 
proof  | 
|
524  | 
fix k  | 
|
525  | 
from fin_f fin_g fin_h fin_fg''  | 
|
526  | 
have "?lhs k = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b when ab = a + b) * h c when k = ab + c)"  | 
|
527  | 
by (simp add: prod_fun_unfold_prod)  | 
|
528  | 
also have "\<dots> = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b * h c when k = ab + c when ab = a + b))"  | 
|
529  | 
apply (subst Sum_any_left_distrib)  | 
|
530  | 
using fin_fg apply (simp add: split_def)  | 
|
531  | 
apply (subst Sum_any_when_independent [symmetric])  | 
|
532  | 
apply (simp add: when_when when_mult mult_when split_def conj_commute)  | 
|
533  | 
done  | 
|
534  | 
also have "\<dots> = (\<Sum>(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)"  | 
|
535  | 
apply (subst Sum_any.cartesian_product2 [of "(?FG' \<times> ?H) \<times> ?FG"])  | 
|
536  | 
apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero)  | 
|
537  | 
done  | 
|
538  | 
also have "\<dots> = (\<Sum>(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)"  | 
|
539  | 
by (rule Sum_any.cong) (simp add: split_def when_def)  | 
|
540  | 
also have "\<dots> = (\<Sum>(ab, cab). (case cab of (c, a, b) \<Rightarrow> f a * g b * h c when k = a + b + c)  | 
|
541  | 
when ab = (case cab of (c, a, b) \<Rightarrow> a + b))"  | 
|
542  | 
by (simp add: split_def)  | 
|
543  | 
also have "\<dots> = (\<Sum>(c, a, b). f a * g b * h c when k = a + b + c)"  | 
|
544  | 
by (simp add: Sum_any_when_dependent_prod_left)  | 
|
545  | 
also have "\<dots> = (\<Sum>(bc, cab). (case cab of (c, a, b) \<Rightarrow> f a * g b * h c when k = a + b + c)  | 
|
546  | 
when bc = (case cab of (c, a, b) \<Rightarrow> b + c))"  | 
|
547  | 
by (simp add: Sum_any_when_dependent_prod_left)  | 
|
548  | 
also have "\<dots> = (\<Sum>(bc, c, a, b). f a * g b * h c when k = a + b + c when bc = b + c)"  | 
|
549  | 
by (simp add: split_def)  | 
|
550  | 
also have "\<dots> = (\<Sum>(bc, c, a, b). f a * g b * h c when bc = b + c when k = a + bc)"  | 
|
551  | 
by (rule Sum_any.cong) (simp add: split_def when_def ac_simps)  | 
|
552  | 
also have "\<dots> = (\<Sum>(a, bc, b, c). f a * g b * h c when bc = b + c when k = a + bc)"  | 
|
553  | 
proof -  | 
|
554  | 
have "bij (\<lambda>(a, d, b, c). (d, c, a, b))"  | 
|
555  | 
by (auto intro!: bijI injI surjI [of _ "\<lambda>(d, c, a, b). (a, d, b, c)"] simp add: split_def)  | 
|
556  | 
then show ?thesis  | 
|
557  | 
by (rule Sum_any.reindex_cong) auto  | 
|
558  | 
qed  | 
|
559  | 
also have "\<dots> = (\<Sum>(a, bc). (\<Sum>(b, c). f a * g b * h c when bc = b + c when k = a + bc))"  | 
|
560  | 
apply (subst Sum_any.cartesian_product2 [of "(?F \<times> ?GH') \<times> ?GH"])  | 
|
561  | 
apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero)  | 
|
562  | 
done  | 
|
563  | 
also have "\<dots> = (\<Sum>(a, bc). f a * (\<Sum>(b, c). g b * h c when bc = b + c) when k = a + bc)"  | 
|
564  | 
apply (subst Sum_any_right_distrib)  | 
|
565  | 
using fin_gh apply (simp add: split_def)  | 
|
566  | 
apply (subst Sum_any_when_independent [symmetric])  | 
|
567  | 
apply (simp add: when_when when_mult mult_when split_def ac_simps)  | 
|
568  | 
done  | 
|
569  | 
also from fin_f fin_g fin_h fin_gh''  | 
|
570  | 
have "\<dots> = ?rhs k"  | 
|
571  | 
by (simp add: prod_fun_unfold_prod)  | 
|
572  | 
finally show "?lhs k = ?rhs k" .  | 
|
573  | 
qed  | 
|
574  | 
qed  | 
|
575  | 
show "(a + b) * c = a * c + b * c"  | 
|
576  | 
proof transfer  | 
|
577  | 
fix f g h :: "'a \<Rightarrow> 'b"  | 
|
578  | 
    assume fin_f: "finite {k. f k \<noteq> 0}"
 | 
|
579  | 
    assume fin_g: "finite {k. g k \<noteq> 0}"
 | 
|
580  | 
    assume fin_h: "finite {k. h k \<noteq> 0}"
 | 
|
581  | 
show "prod_fun (\<lambda>k. f k + g k) h = (\<lambda>k. prod_fun f h k + prod_fun g h k)"  | 
|
582  | 
apply (rule ext)  | 
|
583  | 
apply (auto simp add: prod_fun_def algebra_simps)  | 
|
584  | 
apply (subst Sum_any.distrib)  | 
|
585  | 
using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI)  | 
|
586  | 
done  | 
|
587  | 
qed  | 
|
588  | 
show "a * (b + c) = a * b + a * c"  | 
|
589  | 
proof transfer  | 
|
590  | 
fix f g h :: "'a \<Rightarrow> 'b"  | 
|
591  | 
    assume fin_f: "finite {k. f k \<noteq> 0}"
 | 
|
592  | 
    assume fin_g: "finite {k. g k \<noteq> 0}"
 | 
|
593  | 
    assume fin_h: "finite {k. h k \<noteq> 0}"
 | 
|
594  | 
show "prod_fun f (\<lambda>k. g k + h k) = (\<lambda>k. prod_fun f g k + prod_fun f h k)"  | 
|
595  | 
apply (rule ext)  | 
|
596  | 
apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib)  | 
|
597  | 
apply (subst Sum_any.distrib)  | 
|
598  | 
apply (simp_all add: algebra_simps)  | 
|
599  | 
apply (auto intro: fin_g fin_h)  | 
|
600  | 
apply (subst Sum_any.distrib)  | 
|
601  | 
apply (simp_all add: algebra_simps)  | 
|
602  | 
using fin_f apply (rule finite_mult_not_eq_zero_rightI)  | 
|
603  | 
using fin_f apply (rule finite_mult_not_eq_zero_rightI)  | 
|
604  | 
done  | 
|
605  | 
qed  | 
|
606  | 
show "0 * a = 0"  | 
|
607  | 
by transfer (simp add: prod_fun_def [abs_def])  | 
|
608  | 
show "a * 0 = 0"  | 
|
609  | 
by transfer (simp add: prod_fun_def [abs_def])  | 
|
610  | 
qed  | 
|
611  | 
||
612  | 
end  | 
|
613  | 
||
614  | 
lemma lookup_mult:  | 
|
615  | 
"lookup (f * g) k = (\<Sum>l. lookup f l * (\<Sum>q. lookup g q when k = l + q))"  | 
|
616  | 
by transfer (simp add: prod_fun_def)  | 
|
617  | 
||
618  | 
instance poly_mapping :: (comm_monoid_add, comm_semiring_0) comm_semiring_0  | 
|
619  | 
proof  | 
|
620  | 
fix a b c :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
621  | 
show "a * b = b * a"  | 
|
622  | 
proof transfer  | 
|
623  | 
fix f g :: "'a \<Rightarrow> 'b"  | 
|
624  | 
    assume fin_f: "finite {a. f a \<noteq> 0}"
 | 
|
625  | 
    assume fin_g: "finite {b. g b \<noteq> 0}"
 | 
|
626  | 
show "prod_fun f g = prod_fun g f"  | 
|
627  | 
proof  | 
|
628  | 
fix k  | 
|
629  | 
      have fin1: "\<And>l. finite {a. (f a when k = l + a) \<noteq> 0}"
 | 
|
630  | 
using fin_f by auto  | 
|
631  | 
      have fin2: "\<And>l. finite {b. (g b when k = l + b) \<noteq> 0}"
 | 
|
632  | 
using fin_g by auto  | 
|
633  | 
      from fin_f fin_g have "finite {(a, b). f a \<noteq> 0 \<and> g b \<noteq> 0}" (is "finite ?AB")
 | 
|
634  | 
by simp  | 
|
635  | 
show "prod_fun f g k = prod_fun g f k"  | 
|
636  | 
apply (simp add: prod_fun_def)  | 
|
637  | 
apply (subst Sum_any_right_distrib)  | 
|
638  | 
apply (rule fin2)  | 
|
639  | 
apply (subst Sum_any_right_distrib)  | 
|
640  | 
apply (rule fin1)  | 
|
641  | 
apply (subst Sum_any.swap [of ?AB])  | 
|
642  | 
apply (fact \<open>finite ?AB\<close>)  | 
|
643  | 
apply (auto simp add: mult_when ac_simps)  | 
|
644  | 
done  | 
|
645  | 
qed  | 
|
646  | 
qed  | 
|
647  | 
show "(a + b) * c = a * c + b * c"  | 
|
648  | 
proof transfer  | 
|
649  | 
fix f g h :: "'a \<Rightarrow> 'b"  | 
|
650  | 
    assume fin_f: "finite {k. f k \<noteq> 0}"
 | 
|
651  | 
    assume fin_g: "finite {k. g k \<noteq> 0}"
 | 
|
652  | 
    assume fin_h: "finite {k. h k \<noteq> 0}"
 | 
|
653  | 
show "prod_fun (\<lambda>k. f k + g k) h = (\<lambda>k. prod_fun f h k + prod_fun g h k)"  | 
|
654  | 
apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps)  | 
|
655  | 
apply (subst Sum_any.distrib)  | 
|
656  | 
using fin_f apply (rule finite_mult_not_eq_zero_rightI)  | 
|
657  | 
using fin_g apply (rule finite_mult_not_eq_zero_rightI)  | 
|
658  | 
apply simp_all  | 
|
659  | 
done  | 
|
660  | 
qed  | 
|
661  | 
qed  | 
|
662  | 
||
663  | 
instance poly_mapping :: (monoid_add, semiring_0_cancel) semiring_0_cancel  | 
|
664  | 
..  | 
|
665  | 
||
666  | 
instance poly_mapping :: (comm_monoid_add, comm_semiring_0_cancel) comm_semiring_0_cancel  | 
|
667  | 
..  | 
|
668  | 
||
669  | 
instance poly_mapping :: (monoid_add, semiring_1) semiring_1  | 
|
670  | 
proof  | 
|
671  | 
fix a :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
672  | 
show "1 * a = a"  | 
|
673  | 
by transfer (simp add: prod_fun_def [abs_def] when_mult)  | 
|
674  | 
show "a * 1 = a"  | 
|
675  | 
apply transfer  | 
|
676  | 
apply (simp add: prod_fun_def [abs_def] Sum_any_right_distrib Sum_any_left_distrib mult_when)  | 
|
677  | 
apply (subst when_commute)  | 
|
678  | 
apply simp  | 
|
679  | 
done  | 
|
680  | 
qed  | 
|
681  | 
||
682  | 
instance poly_mapping :: (comm_monoid_add, comm_semiring_1) comm_semiring_1  | 
|
683  | 
proof  | 
|
684  | 
fix a :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
685  | 
show "1 * a = a"  | 
|
686  | 
by transfer (simp add: prod_fun_def [abs_def])  | 
|
687  | 
qed  | 
|
688  | 
||
689  | 
instance poly_mapping :: (monoid_add, semiring_1_cancel) semiring_1_cancel  | 
|
690  | 
..  | 
|
691  | 
||
692  | 
instance poly_mapping :: (monoid_add, ring) ring  | 
|
693  | 
..  | 
|
694  | 
||
695  | 
instance poly_mapping :: (comm_monoid_add, comm_ring) comm_ring  | 
|
696  | 
..  | 
|
697  | 
||
698  | 
instance poly_mapping :: (monoid_add, ring_1) ring_1  | 
|
699  | 
..  | 
|
700  | 
||
701  | 
instance poly_mapping :: (comm_monoid_add, comm_ring_1) comm_ring_1  | 
|
702  | 
..  | 
|
703  | 
||
704  | 
||
705  | 
subsection \<open>Single-point mappings\<close>  | 
|
706  | 
||
707  | 
lift_definition single :: "'a \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b::zero"  | 
|
708  | 
is "\<lambda>k v k'. (v when k = k')"  | 
|
709  | 
by simp  | 
|
710  | 
||
711  | 
lemma inj_single [iff]:  | 
|
712  | 
"inj (single k)"  | 
|
713  | 
proof (rule injI, transfer)  | 
|
714  | 
fix k :: 'b and a b :: "'a::zero"  | 
|
715  | 
assume "(\<lambda>k'. a when k = k') = (\<lambda>k'. b when k = k')"  | 
|
716  | 
then have "(\<lambda>k'. a when k = k') k = (\<lambda>k'. b when k = k') k"  | 
|
717  | 
by (rule arg_cong)  | 
|
718  | 
then show "a = b" by simp  | 
|
719  | 
qed  | 
|
720  | 
||
721  | 
lemma lookup_single:  | 
|
722  | 
"lookup (single k v) k' = (v when k = k')"  | 
|
723  | 
by transfer rule  | 
|
724  | 
||
725  | 
lemma lookup_single_eq [simp]:  | 
|
726  | 
"lookup (single k v) k = v"  | 
|
727  | 
by transfer simp  | 
|
728  | 
||
729  | 
lemma lookup_single_not_eq:  | 
|
730  | 
"k \<noteq> k' \<Longrightarrow> lookup (single k v) k' = 0"  | 
|
731  | 
by transfer simp  | 
|
732  | 
||
733  | 
lemma single_zero [simp]:  | 
|
734  | 
"single k 0 = 0"  | 
|
735  | 
by transfer simp  | 
|
736  | 
||
737  | 
lemma single_one [simp]:  | 
|
738  | 
"single 0 1 = 1"  | 
|
739  | 
by transfer simp  | 
|
740  | 
||
741  | 
lemma single_add:  | 
|
742  | 
"single k (a + b) = single k a + single k b"  | 
|
743  | 
by transfer (simp add: fun_eq_iff when_add_distrib)  | 
|
744  | 
||
745  | 
lemma single_uminus:  | 
|
746  | 
"single k (- a) = - single k a"  | 
|
747  | 
by transfer (simp add: fun_eq_iff when_uminus_distrib)  | 
|
748  | 
||
749  | 
lemma single_diff:  | 
|
750  | 
"single k (a - b) = single k a - single k b"  | 
|
751  | 
by transfer (simp add: fun_eq_iff when_diff_distrib)  | 
|
752  | 
||
753  | 
lemma single_numeral [simp]:  | 
|
754  | 
"single 0 (numeral n) = numeral n"  | 
|
755  | 
by (induct n) (simp_all only: numeral.simps numeral_add single_zero single_one single_add)  | 
|
756  | 
||
757  | 
lemma lookup_numeral:  | 
|
758  | 
"lookup (numeral n) k = (numeral n when k = 0)"  | 
|
759  | 
proof -  | 
|
760  | 
have "lookup (numeral n) k = lookup (single 0 (numeral n)) k"  | 
|
761  | 
by simp  | 
|
762  | 
then show ?thesis unfolding lookup_single by simp  | 
|
763  | 
qed  | 
|
764  | 
||
765  | 
lemma single_of_nat [simp]:  | 
|
766  | 
"single 0 (of_nat n) = of_nat n"  | 
|
767  | 
by (induct n) (simp_all add: single_add)  | 
|
768  | 
||
769  | 
lemma lookup_of_nat:  | 
|
770  | 
"lookup (of_nat n) k = (of_nat n when k = 0)"  | 
|
771  | 
proof -  | 
|
772  | 
have "lookup (of_nat n) k = lookup (single 0 (of_nat n)) k"  | 
|
773  | 
by simp  | 
|
774  | 
then show ?thesis unfolding lookup_single by simp  | 
|
775  | 
qed  | 
|
776  | 
||
777  | 
lemma of_nat_single:  | 
|
778  | 
"of_nat = single 0 \<circ> of_nat"  | 
|
779  | 
by (simp add: fun_eq_iff)  | 
|
780  | 
||
781  | 
lemma mult_single:  | 
|
782  | 
"single k a * single l b = single (k + l) (a * b)"  | 
|
783  | 
proof transfer  | 
|
784  | 
fix k l :: 'a and a b :: 'b  | 
|
785  | 
show "prod_fun (\<lambda>k'. a when k = k') (\<lambda>k'. b when l = k') = (\<lambda>k'. a * b when k + l = k')"  | 
|
786  | 
proof  | 
|
787  | 
fix k'  | 
|
788  | 
have "prod_fun (\<lambda>k'. a when k = k') (\<lambda>k'. b when l = k') k' = (\<Sum>n. a * b when l = n when k' = k + n)"  | 
|
789  | 
by (simp add: prod_fun_def Sum_any_right_distrib mult_when when_mult)  | 
|
790  | 
also have "\<dots> = (\<Sum>n. a * b when k' = k + n when l = n)"  | 
|
791  | 
by (simp add: when_when conj_commute)  | 
|
792  | 
also have "\<dots> = (a * b when k' = k + l)"  | 
|
793  | 
by simp  | 
|
794  | 
also have "\<dots> = (a * b when k + l = k')"  | 
|
795  | 
by (simp add: when_def)  | 
|
796  | 
finally show "prod_fun (\<lambda>k'. a when k = k') (\<lambda>k'. b when l = k') k' =  | 
|
797  | 
(\<lambda>k'. a * b when k + l = k') k'" .  | 
|
798  | 
qed  | 
|
799  | 
qed  | 
|
800  | 
||
801  | 
instance poly_mapping :: (monoid_add, semiring_char_0) semiring_char_0  | 
|
802  | 
by intro_classes (auto intro: inj_compose inj_of_nat simp add: of_nat_single)  | 
|
803  | 
||
804  | 
instance poly_mapping :: (monoid_add, ring_char_0) ring_char_0  | 
|
805  | 
..  | 
|
806  | 
||
807  | 
lemma single_of_int [simp]:  | 
|
808  | 
"single 0 (of_int k) = of_int k"  | 
|
809  | 
by (cases k) (simp_all add: single_diff single_uminus)  | 
|
810  | 
||
811  | 
lemma lookup_of_int:  | 
|
812  | 
"lookup (of_int l) k = (of_int l when k = 0)"  | 
|
813  | 
proof -  | 
|
814  | 
have "lookup (of_int l) k = lookup (single 0 (of_int l)) k"  | 
|
815  | 
by simp  | 
|
816  | 
then show ?thesis unfolding lookup_single by simp  | 
|
817  | 
qed  | 
|
818  | 
||
819  | 
||
820  | 
subsection \<open>Integral domains\<close>  | 
|
821  | 
||
822  | 
instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", semiring_no_zero_divisors) semiring_no_zero_divisors
 | 
|
823  | 
  text \<open>The @{class "linorder"} constraint is a pragmatic device for the proof --- maybe it can be dropped\<close>
 | 
|
824  | 
proof  | 
|
825  | 
fix f g :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
826  | 
assume "f \<noteq> 0" and "g \<noteq> 0"  | 
|
827  | 
then show "f * g \<noteq> 0"  | 
|
828  | 
proof transfer  | 
|
829  | 
fix f g :: "'a \<Rightarrow> 'b"  | 
|
830  | 
    define F where "F = {a. f a \<noteq> 0}"
 | 
|
831  | 
    moreover define G where "G = {a. g a \<noteq> 0}"
 | 
|
832  | 
ultimately have [simp]:  | 
|
833  | 
"\<And>a. f a \<noteq> 0 \<longleftrightarrow> a \<in> F"  | 
|
834  | 
"\<And>b. g b \<noteq> 0 \<longleftrightarrow> b \<in> G"  | 
|
835  | 
by simp_all  | 
|
836  | 
    assume "finite {a. f a \<noteq> 0}"
 | 
|
837  | 
then have [simp]: "finite F"  | 
|
838  | 
by simp  | 
|
839  | 
    assume "finite {a. g a \<noteq> 0}"
 | 
|
840  | 
then have [simp]: "finite G"  | 
|
841  | 
by simp  | 
|
842  | 
assume "f \<noteq> (\<lambda>a. 0)"  | 
|
843  | 
then obtain a where "f a \<noteq> 0"  | 
|
844  | 
by (auto simp add: fun_eq_iff)  | 
|
845  | 
assume "g \<noteq> (\<lambda>b. 0)"  | 
|
846  | 
then obtain b where "g b \<noteq> 0"  | 
|
847  | 
by (auto simp add: fun_eq_iff)  | 
|
848  | 
    from \<open>f a \<noteq> 0\<close> and \<open>g b \<noteq> 0\<close> have "F \<noteq> {}" and "G \<noteq> {}"
 | 
|
849  | 
by auto  | 
|
850  | 
    note Max_F = \<open>finite F\<close> \<open>F \<noteq> {}\<close>
 | 
|
851  | 
    note Max_G = \<open>finite G\<close> \<open>G \<noteq> {}\<close>
 | 
|
852  | 
from Max_F and Max_G have [simp]:  | 
|
853  | 
"Max F \<in> F"  | 
|
854  | 
"Max G \<in> G"  | 
|
855  | 
by auto  | 
|
856  | 
from Max_F Max_G have [dest!]:  | 
|
857  | 
"\<And>a. a \<in> F \<Longrightarrow> a \<le> Max F"  | 
|
858  | 
"\<And>b. b \<in> G \<Longrightarrow> b \<le> Max G"  | 
|
859  | 
by auto  | 
|
860  | 
define q where "q = Max F + Max G"  | 
|
861  | 
have "(\<Sum>(a, b). f a * g b when q = a + b) =  | 
|
862  | 
(\<Sum>(a, b). f a * g b when q = a + b when a \<in> F \<and> b \<in> G)"  | 
|
863  | 
by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr)  | 
|
864  | 
also have "\<dots> =  | 
|
865  | 
(\<Sum>(a, b). f a * g b when (Max F, Max G) = (a, b))"  | 
|
866  | 
proof (rule Sum_any.cong)  | 
|
867  | 
fix ab :: "'a \<times> 'a"  | 
|
868  | 
obtain a b where [simp]: "ab = (a, b)"  | 
|
869  | 
by (cases ab) simp_all  | 
|
870  | 
have [dest!]:  | 
|
871  | 
"a \<le> Max F \<Longrightarrow> Max F \<noteq> a \<Longrightarrow> a < Max F"  | 
|
872  | 
"b \<le> Max G \<Longrightarrow> Max G \<noteq> b \<Longrightarrow> b < Max G"  | 
|
873  | 
by auto  | 
|
874  | 
show "(case ab of (a, b) \<Rightarrow> f a * g b when q = a + b when a \<in> F \<and> b \<in> G) =  | 
|
875  | 
(case ab of (a, b) \<Rightarrow> f a * g b when (Max F, Max G) = (a, b))"  | 
|
876  | 
by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"])  | 
|
877  | 
qed  | 
|
878  | 
also have "\<dots> = (\<Sum>ab. (case ab of (a, b) \<Rightarrow> f a * g b) when  | 
|
879  | 
(Max F, Max G) = ab)"  | 
|
880  | 
unfolding split_def when_def by auto  | 
|
881  | 
also have "\<dots> \<noteq> 0"  | 
|
882  | 
by simp  | 
|
883  | 
finally have "prod_fun f g q \<noteq> 0"  | 
|
884  | 
by (simp add: prod_fun_unfold_prod)  | 
|
885  | 
then show "prod_fun f g \<noteq> (\<lambda>k. 0)"  | 
|
886  | 
by (auto simp add: fun_eq_iff)  | 
|
887  | 
qed  | 
|
888  | 
qed  | 
|
889  | 
||
890  | 
instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_no_zero_divisors) ring_no_zero_divisors
 | 
|
891  | 
..  | 
|
892  | 
||
893  | 
instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_1_no_zero_divisors) ring_1_no_zero_divisors
 | 
|
894  | 
..  | 
|
895  | 
||
896  | 
instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", idom) idom
 | 
|
897  | 
..  | 
|
898  | 
||
899  | 
||
900  | 
subsection \<open>Mapping order\<close>  | 
|
901  | 
||
902  | 
instantiation poly_mapping :: (linorder, "{zero, linorder}") linorder
 | 
|
903  | 
begin  | 
|
904  | 
||
905  | 
lift_definition less_poly_mapping :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
 | 
|
906  | 
is less_fun  | 
|
907  | 
.  | 
|
908  | 
||
909  | 
lift_definition less_eq_poly_mapping :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> bool"
 | 
|
910  | 
is "\<lambda>f g. less_fun f g \<or> f = g"  | 
|
911  | 
.  | 
|
912  | 
||
913  | 
instance proof (rule class.Orderings.linorder.of_class.intro)  | 
|
914  | 
show "class.linorder (less_eq :: (_ \<Rightarrow>\<^sub>0 _) \<Rightarrow> _) less"  | 
|
915  | 
proof (rule linorder_strictI, rule order_strictI)  | 
|
916  | 
fix f g h :: "'a \<Rightarrow>\<^sub>0 'b"  | 
|
917  | 
show "f \<le> g \<longleftrightarrow> f < g \<or> f = g"  | 
|
918  | 
by transfer (rule refl)  | 
|
919  | 
show "\<not> f < f"  | 
|
920  | 
by transfer (rule less_fun_irrefl)  | 
|
921  | 
show "f < g \<or> f = g \<or> g < f"  | 
|
922  | 
proof transfer  | 
|
923  | 
fix f g :: "'a \<Rightarrow> 'b"  | 
|
924  | 
      assume "finite {k. f k \<noteq> 0}" and "finite {k. g k \<noteq> 0}"
 | 
|
925  | 
      then have "finite ({k. f k \<noteq> 0} \<union> {k. g k \<noteq> 0})"
 | 
|
926  | 
by simp  | 
|
927  | 
      moreover have "{k. f k \<noteq> g k} \<subseteq> {k. f k \<noteq> 0} \<union> {k. g k \<noteq> 0}"
 | 
|
928  | 
by auto  | 
|
929  | 
      ultimately have "finite {k. f k \<noteq> g k}"
 | 
|
930  | 
by (rule rev_finite_subset)  | 
|
931  | 
then show "less_fun f g \<or> f = g \<or> less_fun g f"  | 
|
932  | 
by (rule less_fun_trichotomy)  | 
|
933  | 
qed  | 
|
934  | 
assume "f < g" then show "\<not> g < f"  | 
|
935  | 
by transfer (rule less_fun_asym)  | 
|
936  | 
note \<open>f < g\<close> moreover assume "g < h"  | 
|
937  | 
ultimately show "f < h"  | 
|
938  | 
by transfer (rule less_fun_trans)  | 
|
939  | 
qed  | 
|
940  | 
qed  | 
|
941  | 
||
942  | 
end  | 
|
943  | 
||
944  | 
instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, linorder}") ordered_ab_semigroup_add
 | 
|
945  | 
proof (intro_classes, transfer)  | 
|
946  | 
fix f g h :: "'a \<Rightarrow> 'b"  | 
|
947  | 
assume *: "less_fun f g \<or> f = g"  | 
|
948  | 
  { assume "less_fun f g"
 | 
|
949  | 
then obtain k where "f k < g k" "(\<And>k'. k' < k \<Longrightarrow> f k' = g k')"  | 
|
950  | 
by (blast elim!: less_funE)  | 
|
951  | 
then have "h k + f k < h k + g k" "(\<And>k'. k' < k \<Longrightarrow> h k' + f k' = h k' + g k')"  | 
|
952  | 
by simp_all  | 
|
953  | 
then have "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k)"  | 
|
954  | 
by (blast intro: less_funI)  | 
|
955  | 
}  | 
|
956  | 
with * show "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k) \<or> (\<lambda>k. h k + f k) = (\<lambda>k. h k + g k)"  | 
|
957  | 
by (auto simp add: fun_eq_iff)  | 
|
958  | 
qed  | 
|
959  | 
||
960  | 
instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add
 | 
|
961  | 
..  | 
|
962  | 
||
963  | 
instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_comm_monoid_add
 | 
|
964  | 
..  | 
|
965  | 
||
966  | 
instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_cancel_comm_monoid_add
 | 
|
967  | 
..  | 
|
968  | 
||
969  | 
instance poly_mapping :: (linorder, linordered_ab_group_add) linordered_ab_group_add  | 
|
970  | 
..  | 
|
971  | 
||
972  | 
text \<open>  | 
|
973  | 
For pragmatism we leave out the final elements in the hierarchy:  | 
|
974  | 
  @{class linordered_ring}, @{class linordered_ring_strict}, @{class linordered_idom};
 | 
|
975  | 
remember that the order instance is a mere technical device, not a deeper algebraic  | 
|
976  | 
property.  | 
|
977  | 
\<close>  | 
|
978  | 
||
979  | 
||
980  | 
subsection \<open>Fundamental mapping notions\<close>  | 
|
981  | 
||
982  | 
lift_definition keys :: "('a \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 'a set"
 | 
|
983  | 
  is "\<lambda>f. {k. f k \<noteq> 0}" .
 | 
|
984  | 
||
985  | 
lift_definition range :: "('a \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 'b set"
 | 
|
986  | 
  is "\<lambda>f :: 'a \<Rightarrow> 'b. Set.range f - {0}" .
 | 
|
987  | 
||
988  | 
lemma finite_keys [simp]:  | 
|
989  | 
"finite (keys f)"  | 
|
990  | 
by transfer  | 
|
991  | 
||
992  | 
lemma not_in_keys_iff_lookup_eq_zero:  | 
|
993  | 
"k \<notin> keys f \<longleftrightarrow> lookup f k = 0"  | 
|
994  | 
by transfer simp  | 
|
995  | 
||
996  | 
lemma lookup_not_eq_zero_eq_in_keys:  | 
|
997  | 
"lookup f k \<noteq> 0 \<longleftrightarrow> k \<in> keys f"  | 
|
998  | 
by transfer simp  | 
|
999  | 
||
1000  | 
lemma lookup_eq_zero_in_keys_contradict [dest]:  | 
|
1001  | 
"lookup f k = 0 \<Longrightarrow> \<not> k \<in> keys f"  | 
|
1002  | 
by (simp add: not_in_keys_iff_lookup_eq_zero)  | 
|
1003  | 
||
1004  | 
lemma finite_range [simp]: "finite (Poly_Mapping.range p)"  | 
|
1005  | 
proof transfer  | 
|
1006  | 
fix f :: "'b \<Rightarrow> 'a"  | 
|
1007  | 
  assume *: "finite {x. f x \<noteq> 0}"
 | 
|
1008  | 
  have "Set.range f - {0} \<subseteq> f ` {x. f x \<noteq> 0}"
 | 
|
1009  | 
by auto  | 
|
1010  | 
  thus "finite (Set.range f - {0})"
 | 
|
1011  | 
by(rule finite_subset)(rule finite_imageI[OF *])  | 
|
1012  | 
qed  | 
|
1013  | 
||
1014  | 
lemma in_keys_lookup_in_range [simp]:  | 
|
1015  | 
"k \<in> keys f \<Longrightarrow> lookup f k \<in> range f"  | 
|
1016  | 
by transfer simp  | 
|
1017  | 
||
1018  | 
lemma in_keys_iff: "x \<in> (keys s) = (lookup s x \<noteq> 0)"  | 
|
1019  | 
by (transfer, simp)  | 
|
1020  | 
||
1021  | 
lemma keys_zero [simp]:  | 
|
1022  | 
  "keys 0 = {}"
 | 
|
1023  | 
by transfer simp  | 
|
1024  | 
||
1025  | 
lemma range_zero [simp]:  | 
|
1026  | 
  "range 0 = {}"
 | 
|
1027  | 
by transfer auto  | 
|
1028  | 
||
1029  | 
lemma keys_add:  | 
|
1030  | 
"keys (f + g) \<subseteq> keys f \<union> keys g"  | 
|
1031  | 
by transfer auto  | 
|
1032  | 
||
1033  | 
lemma keys_one [simp]:  | 
|
1034  | 
  "keys 1 = {0}"
 | 
|
1035  | 
by transfer simp  | 
|
1036  | 
||
1037  | 
lemma range_one [simp]:  | 
|
1038  | 
  "range 1 = {1}"
 | 
|
1039  | 
by transfer (auto simp add: when_def)  | 
|
1040  | 
||
1041  | 
lemma keys_single [simp]:  | 
|
1042  | 
  "keys (single k v) = (if v = 0 then {} else {k})"
 | 
|
1043  | 
by transfer simp  | 
|
1044  | 
||
1045  | 
lemma range_single [simp]:  | 
|
1046  | 
  "range (single k v) = (if v = 0 then {} else {v})"
 | 
|
1047  | 
by transfer (auto simp add: when_def)  | 
|
1048  | 
||
1049  | 
lemma keys_mult:  | 
|
1050  | 
  "keys (f * g) \<subseteq> {a + b | a b. a \<in> keys f \<and> b \<in> keys g}"
 | 
|
1051  | 
apply transfer  | 
|
1052  | 
apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral)  | 
|
1053  | 
apply blast  | 
|
1054  | 
done  | 
|
1055  | 
||
1056  | 
lemma setsum_keys_plus_distrib:  | 
|
1057  | 
assumes hom_0: "\<And>k. f k 0 = 0"  | 
|
1058  | 
and hom_plus: "\<And>k. k \<in> Poly_Mapping.keys p \<union> Poly_Mapping.keys q \<Longrightarrow> f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k) = f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k)"  | 
|
1059  | 
shows  | 
|
1060  | 
"(\<Sum>k\<in>Poly_Mapping.keys (p + q). f k (Poly_Mapping.lookup (p + q) k)) =  | 
|
1061  | 
(\<Sum>k\<in>Poly_Mapping.keys p. f k (Poly_Mapping.lookup p k)) +  | 
|
1062  | 
(\<Sum>k\<in>Poly_Mapping.keys q. f k (Poly_Mapping.lookup q k))"  | 
|
1063  | 
(is "?lhs = ?p + ?q")  | 
|
1064  | 
proof -  | 
|
1065  | 
let ?A = "Poly_Mapping.keys p \<union> Poly_Mapping.keys q"  | 
|
1066  | 
have "?lhs = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))"  | 
|
1067  | 
apply(rule sum.mono_neutral_cong_left)  | 
|
1068  | 
apply(simp_all add: Poly_Mapping.keys_add)  | 
|
1069  | 
apply(transfer fixing: f)  | 
|
1070  | 
apply(auto simp add: hom_0)[1]  | 
|
1071  | 
apply(transfer fixing: f)  | 
|
1072  | 
apply(auto simp add: hom_0)[1]  | 
|
1073  | 
done  | 
|
1074  | 
also have "\<dots> = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))"  | 
|
1075  | 
by(rule sum.cong)(simp_all add: hom_plus)  | 
|
1076  | 
also have "\<dots> = (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup p k)) + (\<Sum>k\<in>?A. f k (Poly_Mapping.lookup q k))"  | 
|
1077  | 
(is "_ = ?p' + ?q'")  | 
|
1078  | 
by(simp add: sum.distrib)  | 
|
1079  | 
also have "?p' = ?p"  | 
|
1080  | 
by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right)  | 
|
1081  | 
also have "?q' = ?q"  | 
|
1082  | 
by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right)  | 
|
1083  | 
finally show ?thesis .  | 
|
1084  | 
qed  | 
|
1085  | 
||
1086  | 
subsection \<open>Degree\<close>  | 
|
1087  | 
||
1088  | 
definition degree :: "(nat \<Rightarrow>\<^sub>0 'a::zero) \<Rightarrow> nat"  | 
|
1089  | 
where  | 
|
1090  | 
"degree f = Max (insert 0 (Suc ` keys f))"  | 
|
1091  | 
||
1092  | 
lemma degree_zero [simp]:  | 
|
1093  | 
"degree 0 = 0"  | 
|
1094  | 
unfolding degree_def by transfer simp  | 
|
1095  | 
||
1096  | 
lemma degree_one [simp]:  | 
|
1097  | 
"degree 1 = 1"  | 
|
1098  | 
unfolding degree_def by transfer simp  | 
|
1099  | 
||
1100  | 
lemma degree_single_zero [simp]:  | 
|
1101  | 
"degree (single k 0) = 0"  | 
|
1102  | 
unfolding degree_def by transfer simp  | 
|
1103  | 
||
1104  | 
lemma degree_single_not_zero [simp]:  | 
|
1105  | 
"v \<noteq> 0 \<Longrightarrow> degree (single k v) = Suc k"  | 
|
1106  | 
unfolding degree_def by transfer simp  | 
|
1107  | 
||
1108  | 
lemma degree_zero_iff [simp]:  | 
|
1109  | 
"degree f = 0 \<longleftrightarrow> f = 0"  | 
|
1110  | 
unfolding degree_def proof transfer  | 
|
1111  | 
fix f :: "nat \<Rightarrow> 'a"  | 
|
1112  | 
  assume "finite {n. f n \<noteq> 0}"
 | 
|
1113  | 
  then have fin: "finite (insert 0 (Suc ` {n. f n \<noteq> 0}))" by auto
 | 
|
1114  | 
  show "Max (insert 0 (Suc ` {n. f n \<noteq> 0})) = 0 \<longleftrightarrow> f = (\<lambda>n. 0)" (is "?P \<longleftrightarrow> ?Q")
 | 
|
1115  | 
proof  | 
|
1116  | 
assume ?P  | 
|
1117  | 
    have "{n. f n \<noteq> 0} = {}"
 | 
|
1118  | 
proof (rule ccontr)  | 
|
1119  | 
      assume "{n. f n \<noteq> 0} \<noteq> {}"
 | 
|
1120  | 
      then obtain n where "n \<in> {n. f n \<noteq> 0}" by blast
 | 
|
1121  | 
      then have "{n. f n \<noteq> 0} = insert n {n. f n \<noteq> 0}" by auto
 | 
|
1122  | 
      then have "Suc ` {n. f n \<noteq> 0} = insert (Suc n) (Suc ` {n. f n \<noteq> 0})" by auto
 | 
|
1123  | 
      with \<open>?P\<close> have "Max (insert 0 (insert (Suc n) (Suc ` {n. f n \<noteq> 0}))) = 0" by simp
 | 
|
1124  | 
      then have "Max (insert (Suc n) (insert 0 (Suc ` {n. f n \<noteq> 0}))) = 0"
 | 
|
1125  | 
by (simp add: insert_commute)  | 
|
1126  | 
      with fin have "max (Suc n) (Max (insert 0 (Suc ` {n. f n \<noteq> 0}))) = 0"
 | 
|
1127  | 
by simp  | 
|
1128  | 
then show False by simp  | 
|
1129  | 
qed  | 
|
1130  | 
then show ?Q by (simp add: fun_eq_iff)  | 
|
1131  | 
next  | 
|
1132  | 
assume ?Q then show ?P by simp  | 
|
1133  | 
qed  | 
|
1134  | 
qed  | 
|
1135  | 
||
1136  | 
lemma degree_greater_zero_in_keys:  | 
|
1137  | 
assumes "0 < degree f"  | 
|
1138  | 
shows "degree f - 1 \<in> keys f"  | 
|
1139  | 
proof -  | 
|
1140  | 
  from assms have "keys f \<noteq> {}"
 | 
|
1141  | 
by (auto simp add: degree_def)  | 
|
1142  | 
then show ?thesis unfolding degree_def  | 
|
1143  | 
by (simp add: mono_Max_commute [symmetric] mono_Suc)  | 
|
1144  | 
qed  | 
|
1145  | 
||
1146  | 
lemma in_keys_less_degree:  | 
|
1147  | 
"n \<in> keys f \<Longrightarrow> n < degree f"  | 
|
1148  | 
unfolding degree_def by transfer (auto simp add: Max_gr_iff)  | 
|
1149  | 
||
1150  | 
lemma beyond_degree_lookup_zero:  | 
|
1151  | 
"degree f \<le> n \<Longrightarrow> lookup f n = 0"  | 
|
1152  | 
unfolding degree_def by transfer auto  | 
|
1153  | 
||
1154  | 
lemma degree_add:  | 
|
1155  | 
"degree (f + g) \<le> max (degree f) (Poly_Mapping.degree g)"  | 
|
1156  | 
unfolding degree_def proof transfer  | 
|
1157  | 
fix f g :: "nat \<Rightarrow> 'a"  | 
|
1158  | 
  assume f: "finite {x. f x \<noteq> 0}"
 | 
|
1159  | 
  assume g: "finite {x. g x \<noteq> 0}"
 | 
|
1160  | 
  let ?f = "Max (insert 0 (Suc ` {k. f k \<noteq> 0}))"
 | 
|
1161  | 
  let ?g = "Max (insert 0 (Suc ` {k. g k \<noteq> 0}))"
 | 
|
1162  | 
  have "Max (insert 0 (Suc ` {k. f k + g k \<noteq> 0})) \<le> Max (insert 0 (Suc ` ({k. f k \<noteq> 0} \<union> {k. g k \<noteq> 0})))"
 | 
|
1163  | 
by (rule Max.subset_imp) (insert f g, auto)  | 
|
1164  | 
also have "\<dots> = max ?f ?g"  | 
|
1165  | 
using f g by (simp_all add: image_Un Max_Un [symmetric])  | 
|
1166  | 
  finally show "Max (insert 0 (Suc ` {k. f k + g k \<noteq> 0}))
 | 
|
1167  | 
    \<le> max (Max (insert 0 (Suc ` {k. f k \<noteq> 0}))) (Max (insert 0 (Suc ` {k. g k \<noteq> 0})))"
 | 
|
1168  | 
.  | 
|
1169  | 
qed  | 
|
1170  | 
||
1171  | 
lemma sorted_list_of_set_keys:  | 
|
1172  | 
"sorted_list_of_set (keys f) = filter (\<lambda>k. k \<in> keys f) [0..<degree f]" (is "_ = ?r")  | 
|
1173  | 
proof -  | 
|
1174  | 
have "keys f = set ?r"  | 
|
1175  | 
by (auto dest: in_keys_less_degree)  | 
|
1176  | 
moreover have "sorted_list_of_set (set ?r) = ?r"  | 
|
1177  | 
unfolding sorted_list_of_set_sort_remdups  | 
|
1178  | 
by (simp add: remdups_filter filter_sort [symmetric])  | 
|
1179  | 
ultimately show ?thesis by simp  | 
|
1180  | 
qed  | 
|
1181  | 
||
1182  | 
||
1183  | 
subsection \<open>Inductive structure\<close>  | 
|
1184  | 
||
1185  | 
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b"
 | 
|
1186  | 
is "\<lambda>k v f. f(k := v)"  | 
|
1187  | 
proof -  | 
|
1188  | 
fix f :: "'a \<Rightarrow> 'b" and k' v  | 
|
1189  | 
  assume "finite {k. f k \<noteq> 0}"
 | 
|
1190  | 
  then have "finite (insert k' {k. f k \<noteq> 0})"
 | 
|
1191  | 
by simp  | 
|
1192  | 
  then show "finite {k. (f(k' := v)) k \<noteq> 0}"
 | 
|
1193  | 
by (rule rev_finite_subset) auto  | 
|
1194  | 
qed  | 
|
1195  | 
||
1196  | 
lemma update_induct [case_names const update]:  | 
|
1197  | 
assumes const': "P 0"  | 
|
1198  | 
assumes update': "\<And>f a b. a \<notin> keys f \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> P f \<Longrightarrow> P (update a b f)"  | 
|
1199  | 
shows "P f"  | 
|
1200  | 
proof -  | 
|
1201  | 
  obtain g where "f = Abs_poly_mapping g" and "finite {a. g a \<noteq> 0}"
 | 
|
1202  | 
by (cases f) simp_all  | 
|
1203  | 
define Q where "Q g = P (Abs_poly_mapping g)" for g  | 
|
1204  | 
  from \<open>finite {a. g a \<noteq> 0}\<close> have "Q g"
 | 
|
1205  | 
proof (induct g rule: finite_update_induct)  | 
|
1206  | 
case const with const' Q_def show ?case  | 
|
1207  | 
by simp  | 
|
1208  | 
next  | 
|
1209  | 
case (update a b g)  | 
|
1210  | 
    from \<open>finite {a. g a \<noteq> 0}\<close> \<open>g a = 0\<close> have "a \<notin> keys (Abs_poly_mapping g)"
 | 
|
1211  | 
by (simp add: Abs_poly_mapping_inverse keys.rep_eq)  | 
|
1212  | 
moreover note \<open>b \<noteq> 0\<close>  | 
|
1213  | 
moreover from \<open>Q g\<close> have "P (Abs_poly_mapping g)"  | 
|
1214  | 
by (simp add: Q_def)  | 
|
1215  | 
ultimately have "P (update a b (Abs_poly_mapping g))"  | 
|
1216  | 
by (rule update')  | 
|
1217  | 
    also from \<open>finite {a. g a \<noteq> 0}\<close>
 | 
|
1218  | 
have "update a b (Abs_poly_mapping g) = Abs_poly_mapping (g(a := b))"  | 
|
1219  | 
by (simp add: update.abs_eq eq_onp_same_args)  | 
|
1220  | 
finally show ?case  | 
|
1221  | 
by (simp add: Q_def fun_upd_def)  | 
|
1222  | 
qed  | 
|
1223  | 
then show ?thesis by (simp add: Q_def \<open>f = Abs_poly_mapping g\<close>)  | 
|
1224  | 
qed  | 
|
1225  | 
||
1226  | 
lemma lookup_update:  | 
|
1227  | 
"lookup (update k v f) k' = (if k = k' then v else lookup f k')"  | 
|
1228  | 
by transfer simp  | 
|
1229  | 
||
1230  | 
lemma keys_update:  | 
|
1231  | 
  "keys (update k v f) = (if v = 0 then keys f - {k} else insert k (keys f))"
 | 
|
1232  | 
by transfer auto  | 
|
1233  | 
||
1234  | 
||
1235  | 
subsection \<open>Quasi-functorial structure\<close>  | 
|
1236  | 
||
1237  | 
lift_definition map :: "('b::zero \<Rightarrow> 'c::zero)
 | 
|
1238  | 
  \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'c::zero)"
 | 
|
1239  | 
is "\<lambda>g f k. g (f k) when f k \<noteq> 0"  | 
|
1240  | 
by simp  | 
|
1241  | 
||
1242  | 
context  | 
|
1243  | 
fixes f :: "'b \<Rightarrow> 'a"  | 
|
1244  | 
assumes inj_f: "inj f"  | 
|
1245  | 
begin  | 
|
1246  | 
||
1247  | 
lift_definition map_key :: "('a \<Rightarrow>\<^sub>0 'c::zero) \<Rightarrow> 'b \<Rightarrow>\<^sub>0 'c"
 | 
|
1248  | 
is "\<lambda>p. p \<circ> f"  | 
|
1249  | 
proof -  | 
|
1250  | 
fix g :: "'c \<Rightarrow> 'd" and p :: "'a \<Rightarrow> 'c"  | 
|
1251  | 
  assume "finite {x. p x \<noteq> 0}"
 | 
|
1252  | 
  hence "finite (f ` {y. p (f y) \<noteq> 0})"
 | 
|
1253  | 
by(rule finite_subset[rotated]) auto  | 
|
1254  | 
  thus "finite {x. (p \<circ> f) x \<noteq> 0}" unfolding o_def
 | 
|
1255  | 
by(rule finite_imageD)(rule subset_inj_on[OF inj_f], simp)  | 
|
1256  | 
qed  | 
|
1257  | 
||
1258  | 
end  | 
|
1259  | 
||
1260  | 
lemma map_key_compose:  | 
|
1261  | 
assumes [transfer_rule]: "inj f" "inj g"  | 
|
1262  | 
shows "map_key f (map_key g p) = map_key (g \<circ> f) p"  | 
|
1263  | 
proof -  | 
|
1264  | 
from assms have [transfer_rule]: "inj (g \<circ> f)"  | 
|
1265  | 
by(simp add: inj_compose)  | 
|
1266  | 
show ?thesis by transfer(simp add: o_assoc)  | 
|
1267  | 
qed  | 
|
1268  | 
||
1269  | 
lemma map_key_id:  | 
|
1270  | 
"map_key (\<lambda>x. x) p = p"  | 
|
1271  | 
proof -  | 
|
1272  | 
have [transfer_rule]: "inj (\<lambda>x. x)" by simp  | 
|
1273  | 
show ?thesis by transfer(simp add: o_def)  | 
|
1274  | 
qed  | 
|
1275  | 
||
1276  | 
context  | 
|
1277  | 
fixes f :: "'a \<Rightarrow> 'b"  | 
|
1278  | 
assumes inj_f [transfer_rule]: "inj f"  | 
|
1279  | 
begin  | 
|
1280  | 
||
1281  | 
lemma map_key_map:  | 
|
1282  | 
"map_key f (map g p) = map g (map_key f p)"  | 
|
1283  | 
by transfer (simp add: fun_eq_iff)  | 
|
1284  | 
||
1285  | 
lemma map_key_plus:  | 
|
1286  | 
"map_key f (p + q) = map_key f p + map_key f q"  | 
|
1287  | 
by transfer (simp add: fun_eq_iff)  | 
|
1288  | 
||
1289  | 
lemma keys_map_key:  | 
|
1290  | 
"keys (map_key f p) = f -` keys p"  | 
|
1291  | 
by transfer auto  | 
|
1292  | 
||
1293  | 
lemma map_key_zero [simp]:  | 
|
1294  | 
"map_key f 0 = 0"  | 
|
1295  | 
by transfer (simp add: fun_eq_iff)  | 
|
1296  | 
||
1297  | 
lemma map_key_single [simp]:  | 
|
1298  | 
"map_key f (single (f k) v) = single k v"  | 
|
1299  | 
by transfer (simp add: fun_eq_iff inj_onD [OF inj_f] when_def)  | 
|
1300  | 
||
1301  | 
end  | 
|
1302  | 
||
1303  | 
lemma mult_map_scale_conv_mult: "map ((*) s) p = single 0 s * p"  | 
|
1304  | 
proof(transfer fixing: s)  | 
|
1305  | 
fix p :: "'a \<Rightarrow> 'b"  | 
|
1306  | 
  assume *: "finite {x. p x \<noteq> 0}"
 | 
|
1307  | 
  { fix x
 | 
|
1308  | 
have "prod_fun (\<lambda>k'. s when 0 = k') p x =  | 
|
1309  | 
(\<Sum>l :: 'a. if l = 0 then s * (\<Sum>q. p q when x = q) else 0)"  | 
|
1310  | 
by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta)  | 
|
1311  | 
also have "\<dots> = (\<lambda>k. s * p k when p k \<noteq> 0) x" by(simp add: when_def)  | 
|
1312  | 
also note calculation }  | 
|
1313  | 
then show "(\<lambda>k. s * p k when p k \<noteq> 0) = prod_fun (\<lambda>k'. s when 0 = k') p"  | 
|
1314  | 
by(simp add: fun_eq_iff)  | 
|
1315  | 
qed  | 
|
1316  | 
||
1317  | 
lemma map_single [simp]:  | 
|
1318  | 
"(c = 0 \<Longrightarrow> f 0 = 0) \<Longrightarrow> map f (single x c) = single x (f c)"  | 
|
1319  | 
by transfer(auto simp add: fun_eq_iff when_def)  | 
|
1320  | 
||
1321  | 
lemma map_eq_zero_iff: "map f p = 0 \<longleftrightarrow> (\<forall>k \<in> keys p. f (lookup p k) = 0)"  | 
|
1322  | 
by transfer(auto simp add: fun_eq_iff when_def)  | 
|
1323  | 
||
1324  | 
subsection \<open>Canonical dense representation of @{typ "nat \<Rightarrow>\<^sub>0 'a"}\<close>
 | 
|
1325  | 
||
1326  | 
abbreviation no_trailing_zeros :: "'a :: zero list \<Rightarrow> bool"  | 
|
1327  | 
where  | 
|
1328  | 
"no_trailing_zeros \<equiv> no_trailing ((=) 0)"  | 
|
1329  | 
||
1330  | 
lift_definition "nth" :: "'a list \<Rightarrow> (nat \<Rightarrow>\<^sub>0 'a::zero)"  | 
|
1331  | 
is "nth_default 0"  | 
|
1332  | 
by (fact finite_nth_default_neq_default)  | 
|
1333  | 
||
1334  | 
text \<open>  | 
|
1335  | 
The opposite direction is directly specified on (later)  | 
|
1336  | 
type \<open>nat_mapping\<close>.  | 
|
1337  | 
\<close>  | 
|
1338  | 
||
1339  | 
lemma nth_Nil [simp]:  | 
|
1340  | 
"nth [] = 0"  | 
|
1341  | 
by transfer (simp add: fun_eq_iff)  | 
|
1342  | 
||
1343  | 
lemma nth_singleton [simp]:  | 
|
1344  | 
"nth [v] = single 0 v"  | 
|
1345  | 
proof (transfer, rule ext)  | 
|
1346  | 
fix n :: nat and v :: 'a  | 
|
1347  | 
show "nth_default 0 [v] n = (v when 0 = n)"  | 
|
1348  | 
by (auto simp add: nth_default_def nth_append)  | 
|
1349  | 
qed  | 
|
1350  | 
||
1351  | 
lemma nth_replicate [simp]:  | 
|
1352  | 
"nth (replicate n 0 @ [v]) = single n v"  | 
|
1353  | 
proof (transfer, rule ext)  | 
|
1354  | 
fix m n :: nat and v :: 'a  | 
|
1355  | 
show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)"  | 
|
1356  | 
by (auto simp add: nth_default_def nth_append)  | 
|
1357  | 
qed  | 
|
1358  | 
||
1359  | 
lemma nth_strip_while [simp]:  | 
|
1360  | 
"nth (strip_while ((=) 0) xs) = nth xs"  | 
|
1361  | 
by transfer (fact nth_default_strip_while_dflt)  | 
|
1362  | 
||
1363  | 
lemma nth_strip_while' [simp]:  | 
|
1364  | 
"nth (strip_while (\<lambda>k. k = 0) xs) = nth xs"  | 
|
1365  | 
by (subst eq_commute) (fact nth_strip_while)  | 
|
1366  | 
||
1367  | 
lemma nth_eq_iff:  | 
|
1368  | 
"nth xs = nth ys \<longleftrightarrow> strip_while (HOL.eq 0) xs = strip_while (HOL.eq 0) ys"  | 
|
1369  | 
by transfer (simp add: nth_default_eq_iff)  | 
|
1370  | 
||
1371  | 
lemma lookup_nth [simp]:  | 
|
1372  | 
"lookup (nth xs) = nth_default 0 xs"  | 
|
1373  | 
by (fact nth.rep_eq)  | 
|
1374  | 
||
1375  | 
lemma keys_nth [simp]:  | 
|
1376  | 
  "keys (nth xs) =  fst ` {(n, v) \<in> set (enumerate 0 xs). v \<noteq> 0}"
 | 
|
1377  | 
proof transfer  | 
|
1378  | 
fix xs :: "'a list"  | 
|
1379  | 
  { fix n
 | 
|
1380  | 
assume "nth_default 0 xs n \<noteq> 0"  | 
|
1381  | 
then have "n < length xs" and "xs ! n \<noteq> 0"  | 
|
1382  | 
by (auto simp add: nth_default_def split: if_splits)  | 
|
1383  | 
    then have "(n, xs ! n) \<in> {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}" (is "?x \<in> ?A")
 | 
|
1384  | 
by (auto simp add: in_set_conv_nth enumerate_eq_zip)  | 
|
1385  | 
then have "fst ?x \<in> fst ` ?A"  | 
|
1386  | 
by blast  | 
|
1387  | 
    then have "n \<in> fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
 | 
|
1388  | 
by simp  | 
|
1389  | 
}  | 
|
1390  | 
  then show "{k. nth_default 0 xs k \<noteq> 0} = fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
 | 
|
1391  | 
by (auto simp add: in_enumerate_iff_nth_default_eq)  | 
|
1392  | 
qed  | 
|
1393  | 
||
1394  | 
lemma range_nth [simp]:  | 
|
1395  | 
  "range (nth xs) = set xs - {0}"
 | 
|
1396  | 
by transfer simp  | 
|
1397  | 
||
1398  | 
lemma degree_nth:  | 
|
1399  | 
"no_trailing_zeros xs \<Longrightarrow> degree (nth xs) = length xs"  | 
|
1400  | 
unfolding degree_def proof transfer  | 
|
1401  | 
fix xs :: "'a list"  | 
|
1402  | 
assume *: "no_trailing_zeros xs"  | 
|
1403  | 
  let ?A = "{n. nth_default 0 xs n \<noteq> 0}"
 | 
|
1404  | 
let ?f = "nth_default 0 xs"  | 
|
1405  | 
  let ?bound = "Max (insert 0 (Suc ` {n. ?f n \<noteq> 0}))"
 | 
|
1406  | 
show "?bound = length xs"  | 
|
1407  | 
proof (cases "xs = []")  | 
|
1408  | 
case False  | 
|
1409  | 
with * obtain n where n: "n < length xs" "xs ! n \<noteq> 0"  | 
|
1410  | 
by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv)  | 
|
1411  | 
    then have "?bound = Max (Suc ` {k. (k < length xs \<longrightarrow> xs ! k \<noteq> (0::'a)) \<and> k < length xs})"
 | 
|
1412  | 
by (subst Max_insert) (auto simp add: nth_default_def)  | 
|
1413  | 
    also let ?A = "{k. k < length xs \<and> xs ! k \<noteq> 0}"
 | 
|
1414  | 
    have "{k. (k < length xs \<longrightarrow> xs ! k \<noteq> (0::'a)) \<and> k < length xs} = ?A" by auto
 | 
|
1415  | 
also have "Max (Suc ` ?A) = Suc (Max ?A)" using n  | 
|
1416  | 
by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc)  | 
|
1417  | 
    also {
 | 
|
1418  | 
have "Max ?A \<in> ?A" using n Max_in [of ?A] by fastforce  | 
|
1419  | 
hence "Suc (Max ?A) \<le> length xs" by simp  | 
|
1420  | 
moreover from * False have "length xs - 1 \<in> ?A"  | 
|
1421  | 
by(auto simp add: no_trailing_unfold last_conv_nth)  | 
|
1422  | 
hence "length xs - 1 \<le> Max ?A" using Max_ge[of ?A "length xs - 1"] by auto  | 
|
1423  | 
hence "length xs \<le> Suc (Max ?A)" by simp  | 
|
1424  | 
ultimately have "Suc (Max ?A) = length xs" by simp }  | 
|
1425  | 
finally show ?thesis .  | 
|
1426  | 
qed simp  | 
|
1427  | 
qed  | 
|
1428  | 
||
1429  | 
lemma nth_trailing_zeros [simp]:  | 
|
1430  | 
"nth (xs @ replicate n 0) = nth xs"  | 
|
1431  | 
by transfer simp  | 
|
1432  | 
||
1433  | 
lemma nth_idem:  | 
|
1434  | 
"nth (List.map (lookup f) [0..<degree f]) = f"  | 
|
1435  | 
unfolding degree_def by transfer  | 
|
1436  | 
(auto simp add: nth_default_def fun_eq_iff not_less)  | 
|
1437  | 
||
1438  | 
lemma nth_idem_bound:  | 
|
1439  | 
assumes "degree f \<le> n"  | 
|
1440  | 
shows "nth (List.map (lookup f) [0..<n]) = f"  | 
|
1441  | 
proof -  | 
|
1442  | 
from assms obtain m where "n = degree f + m"  | 
|
1443  | 
by (blast dest: le_Suc_ex)  | 
|
1444  | 
then have "[0..<n] = [0..<degree f] @ [degree f..<degree f + m]"  | 
|
1445  | 
by (simp add: upt_add_eq_append [of 0])  | 
|
1446  | 
moreover have "List.map (lookup f) [degree f..<degree f + m] = replicate m 0"  | 
|
1447  | 
by (rule replicate_eqI) (auto simp add: beyond_degree_lookup_zero)  | 
|
1448  | 
ultimately show ?thesis by (simp add: nth_idem)  | 
|
1449  | 
qed  | 
|
1450  | 
||
1451  | 
||
1452  | 
subsection \<open>Canonical sparse representation of @{typ "'a \<Rightarrow>\<^sub>0 'b"}\<close>
 | 
|
1453  | 
||
1454  | 
lift_definition the_value :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow>\<^sub>0 'b::zero"
 | 
|
1455  | 
is "\<lambda>xs k. case map_of xs k of None \<Rightarrow> 0 | Some v \<Rightarrow> v"  | 
|
1456  | 
proof -  | 
|
1457  | 
  fix xs :: "('a \<times> 'b) list"
 | 
|
1458  | 
  have fin: "finite {k. \<exists>v. map_of xs k = Some v}"
 | 
|
1459  | 
using finite_dom_map_of [of xs] unfolding dom_def by auto  | 
|
1460  | 
  then show "finite {k. (case map_of xs k of None \<Rightarrow> 0 | Some v \<Rightarrow> v) \<noteq> 0}"
 | 
|
1461  | 
using fin by (simp split: option.split)  | 
|
1462  | 
qed  | 
|
1463  | 
||
1464  | 
definition items :: "('a::linorder \<Rightarrow>\<^sub>0 'b::zero) \<Rightarrow> ('a \<times> 'b) list"
 | 
|
1465  | 
where  | 
|
1466  | 
"items f = List.map (\<lambda>k. (k, lookup f k)) (sorted_list_of_set (keys f))"  | 
|
1467  | 
||
1468  | 
text \<open>  | 
|
1469  | 
For the canonical sparse representation we provide both  | 
|
1470  | 
directions of morphisms since the specification of  | 
|
1471  | 
ordered association lists in theory \<open>OAList\<close>  | 
|
1472  | 
  will support arbitrary linear orders @{class linorder}
 | 
|
1473  | 
  as keys, not just natural numbers @{typ nat}.
 | 
|
1474  | 
\<close>  | 
|
1475  | 
||
1476  | 
lemma the_value_items [simp]:  | 
|
1477  | 
"the_value (items f) = f"  | 
|
1478  | 
unfolding items_def  | 
|
1479  | 
by transfer (simp add: fun_eq_iff map_of_map_restrict restrict_map_def)  | 
|
1480  | 
||
1481  | 
lemma lookup_the_value:  | 
|
1482  | 
"lookup (the_value xs) k = (case map_of xs k of None \<Rightarrow> 0 | Some v \<Rightarrow> v)"  | 
|
1483  | 
by transfer rule  | 
|
1484  | 
||
1485  | 
lemma items_the_value:  | 
|
1486  | 
assumes "sorted (List.map fst xs)" and "distinct (List.map fst xs)" and "0 \<notin> snd ` set xs"  | 
|
1487  | 
shows "items (the_value xs) = xs"  | 
|
1488  | 
proof -  | 
|
1489  | 
from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs"  | 
|
1490  | 
unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sorted_sort_id)  | 
|
1491  | 
moreover from assms have "keys (the_value xs) = fst ` set xs"  | 
|
1492  | 
by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr)  | 
|
1493  | 
ultimately show ?thesis  | 
|
1494  | 
unfolding items_def using assms  | 
|
1495  | 
by (auto simp add: lookup_the_value intro: map_idI)  | 
|
1496  | 
qed  | 
|
1497  | 
||
1498  | 
lemma the_value_Nil [simp]:  | 
|
1499  | 
"the_value [] = 0"  | 
|
1500  | 
by transfer (simp add: fun_eq_iff)  | 
|
1501  | 
||
1502  | 
lemma the_value_Cons [simp]:  | 
|
1503  | 
"the_value (x # xs) = update (fst x) (snd x) (the_value xs)"  | 
|
1504  | 
by transfer (simp add: fun_eq_iff)  | 
|
1505  | 
||
1506  | 
lemma items_zero [simp]:  | 
|
1507  | 
"items 0 = []"  | 
|
1508  | 
unfolding items_def by simp  | 
|
1509  | 
||
1510  | 
lemma items_one [simp]:  | 
|
1511  | 
"items 1 = [(0, 1)]"  | 
|
1512  | 
unfolding items_def by transfer simp  | 
|
1513  | 
||
1514  | 
lemma items_single [simp]:  | 
|
1515  | 
"items (single k v) = (if v = 0 then [] else [(k, v)])"  | 
|
1516  | 
unfolding items_def by simp  | 
|
1517  | 
||
1518  | 
lemma in_set_items_iff [simp]:  | 
|
1519  | 
"(k, v) \<in> set (items f) \<longleftrightarrow> k \<in> keys f \<and> lookup f k = v"  | 
|
1520  | 
unfolding items_def by transfer auto  | 
|
1521  | 
||
1522  | 
||
1523  | 
subsection \<open>Size estimation\<close>  | 
|
1524  | 
||
1525  | 
context  | 
|
1526  | 
fixes f :: "'a \<Rightarrow> nat"  | 
|
1527  | 
and g :: "'b :: zero \<Rightarrow> nat"  | 
|
1528  | 
begin  | 
|
1529  | 
||
1530  | 
definition poly_mapping_size :: "('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> nat"
 | 
|
1531  | 
where  | 
|
1532  | 
"poly_mapping_size m = g 0 + (\<Sum>k \<in> keys m. Suc (f k + g (lookup m k)))"  | 
|
1533  | 
||
1534  | 
lemma poly_mapping_size_0 [simp]:  | 
|
1535  | 
"poly_mapping_size 0 = g 0"  | 
|
1536  | 
by (simp add: poly_mapping_size_def)  | 
|
1537  | 
||
1538  | 
lemma poly_mapping_size_single [simp]:  | 
|
1539  | 
"poly_mapping_size (single k v) = (if v = 0 then g 0 else g 0 + f k + g v + 1)"  | 
|
1540  | 
unfolding poly_mapping_size_def by transfer simp  | 
|
1541  | 
||
1542  | 
lemma keys_less_poly_mapping_size:  | 
|
1543  | 
"k \<in> keys m \<Longrightarrow> f k + g (lookup m k) < poly_mapping_size m"  | 
|
1544  | 
unfolding poly_mapping_size_def  | 
|
1545  | 
proof transfer  | 
|
1546  | 
fix k :: 'a and m :: "'a \<Rightarrow> 'b" and f :: "'a \<Rightarrow> nat" and g  | 
|
1547  | 
  let ?keys = "{k. m k \<noteq> 0}"
 | 
|
1548  | 
assume *: "finite ?keys" "k \<in> ?keys"  | 
|
1549  | 
then have "f k + g (m k) = (\<Sum>k' \<in> ?keys. f k' + g (m k') when k' = k)"  | 
|
1550  | 
by (simp add: sum.delta when_def)  | 
|
1551  | 
also have "\<dots> < (\<Sum>k' \<in> ?keys. Suc (f k' + g (m k')))" using *  | 
|
1552  | 
by (intro sum_strict_mono) (auto simp add: when_def)  | 
|
1553  | 
also have "\<dots> \<le> g 0 + \<dots>" by simp  | 
|
1554  | 
finally have "f k + g (m k) < \<dots>" .  | 
|
1555  | 
then show "f k + g (m k) < g 0 + (\<Sum>k | m k \<noteq> 0. Suc (f k + g (m k)))"  | 
|
1556  | 
by simp  | 
|
1557  | 
qed  | 
|
1558  | 
||
1559  | 
lemma lookup_le_poly_mapping_size:  | 
|
1560  | 
"g (lookup m k) \<le> poly_mapping_size m"  | 
|
1561  | 
proof (cases "k \<in> keys m")  | 
|
1562  | 
case True  | 
|
1563  | 
with keys_less_poly_mapping_size [of k m]  | 
|
1564  | 
show ?thesis by simp  | 
|
1565  | 
next  | 
|
1566  | 
case False  | 
|
1567  | 
then show ?thesis  | 
|
1568  | 
by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff)  | 
|
1569  | 
qed  | 
|
1570  | 
||
1571  | 
lemma poly_mapping_size_estimation:  | 
|
1572  | 
"k \<in> keys m \<Longrightarrow> y \<le> f k + g (lookup m k) \<Longrightarrow> y < poly_mapping_size m"  | 
|
1573  | 
using keys_less_poly_mapping_size by (auto intro: le_less_trans)  | 
|
1574  | 
||
1575  | 
lemma poly_mapping_size_estimation2:  | 
|
1576  | 
assumes "v \<in> range m" and "y \<le> g v"  | 
|
1577  | 
shows "y < poly_mapping_size m"  | 
|
1578  | 
proof -  | 
|
1579  | 
from assms obtain k where *: "lookup m k = v" "v \<noteq> 0"  | 
|
1580  | 
by transfer blast  | 
|
1581  | 
from * have "k \<in> keys m"  | 
|
1582  | 
by (simp add: in_keys_iff)  | 
|
1583  | 
then show ?thesis  | 
|
1584  | 
proof (rule poly_mapping_size_estimation)  | 
|
1585  | 
from assms * have "y \<le> g (lookup m k)"  | 
|
1586  | 
by simp  | 
|
1587  | 
then show "y \<le> f k + g (lookup m k)"  | 
|
1588  | 
by simp  | 
|
1589  | 
qed  | 
|
1590  | 
qed  | 
|
1591  | 
||
1592  | 
end  | 
|
1593  | 
||
1594  | 
lemma poly_mapping_size_one [simp]:  | 
|
1595  | 
"poly_mapping_size f g 1 = g 0 + f 0 + g 1 + 1"  | 
|
1596  | 
unfolding poly_mapping_size_def by transfer simp  | 
|
1597  | 
||
1598  | 
lemma poly_mapping_size_cong [fundef_cong]:  | 
|
1599  | 
"m = m' \<Longrightarrow> g 0 = g' 0 \<Longrightarrow> (\<And>k. k \<in> keys m' \<Longrightarrow> f k = f' k)  | 
|
1600  | 
\<Longrightarrow> (\<And>v. v \<in> range m' \<Longrightarrow> g v = g' v)  | 
|
1601  | 
\<Longrightarrow> poly_mapping_size f g m = poly_mapping_size f' g' m'"  | 
|
1602  | 
by (auto simp add: poly_mapping_size_def intro!: sum.cong)  | 
|
1603  | 
||
1604  | 
instantiation poly_mapping :: (type, zero) size  | 
|
1605  | 
begin  | 
|
1606  | 
||
1607  | 
definition "size = poly_mapping_size (\<lambda>_. 0) (\<lambda>_. 0)"  | 
|
1608  | 
||
1609  | 
instance ..  | 
|
1610  | 
||
1611  | 
end  | 
|
1612  | 
||
1613  | 
||
1614  | 
subsection \<open>Further mapping operations and properties\<close>  | 
|
1615  | 
||
1616  | 
text \<open>It is like in algebra: there are many definitions, some are also used\<close>  | 
|
1617  | 
||
1618  | 
lift_definition mapp ::  | 
|
1619  | 
  "('a \<Rightarrow> 'b :: zero \<Rightarrow> 'c :: zero) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'b) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 'c)"
 | 
|
1620  | 
is "\<lambda>f p k. (if k \<in> keys p then f k (lookup p k) else 0)"  | 
|
1621  | 
by simp  | 
|
1622  | 
||
1623  | 
lemma mapp_cong [fundef_cong]:  | 
|
1624  | 
"\<lbrakk> m = m'; \<And>k. k \<in> keys m' \<Longrightarrow> f k (lookup m' k) = f' k (lookup m' k) \<rbrakk>  | 
|
1625  | 
\<Longrightarrow> mapp f m = mapp f' m'"  | 
|
1626  | 
by transfer (auto simp add: fun_eq_iff)  | 
|
1627  | 
||
1628  | 
lemma lookup_mapp:  | 
|
1629  | 
"lookup (mapp f p) k = (f k (lookup p k) when k \<in> keys p)"  | 
|
1630  | 
by (simp add: mapp.rep_eq)  | 
|
1631  | 
||
1632  | 
lemma keys_mapp_subset: "keys (mapp f p) \<subseteq> keys p"  | 
|
1633  | 
by (meson in_keys_iff mapp.rep_eq subsetI)  | 
|
1634  | 
||
1635  | 
subsection\<open>Free Abelian Groups Over a Type\<close>  | 
|
1636  | 
||
1637  | 
abbreviation frag_of :: "'a \<Rightarrow> 'a \<Rightarrow>\<^sub>0 int"  | 
|
1638  | 
where "frag_of c \<equiv> Poly_Mapping.single c (1::int)"  | 
|
1639  | 
||
1640  | 
lemma lookup_frag_of [simp]:  | 
|
1641  | 
"Poly_Mapping.lookup(frag_of c) = (\<lambda>x. if x = c then 1 else 0)"  | 
|
1642  | 
by (force simp add: lookup_single_not_eq)  | 
|
1643  | 
||
1644  | 
lemma frag_of_nonzero [simp]: "frag_of a \<noteq> 0"  | 
|
1645  | 
proof -  | 
|
1646  | 
let ?f = "\<lambda>x. if x = a then 1 else (0::int)"  | 
|
1647  | 
have "?f \<noteq> (\<lambda>x. 0::int)"  | 
|
1648  | 
by (auto simp: fun_eq_iff)  | 
|
1649  | 
then have "Poly_Mapping.lookup (Abs_poly_mapping ?f) \<noteq> Poly_Mapping.lookup (Abs_poly_mapping (\<lambda>x. 0))"  | 
|
1650  | 
by fastforce  | 
|
1651  | 
then show ?thesis  | 
|
1652  | 
by (metis lookup_single_eq lookup_zero)  | 
|
1653  | 
qed  | 
|
1654  | 
||
1655  | 
definition frag_cmul :: "int \<Rightarrow> ('a \<Rightarrow>\<^sub>0 int) \<Rightarrow> ('a \<Rightarrow>\<^sub>0 int)"
 | 
|
1656  | 
where "frag_cmul c a = Abs_poly_mapping (\<lambda>x. c * Poly_Mapping.lookup a x)"  | 
|
1657  | 
||
1658  | 
lemma frag_cmul_zero [simp]: "frag_cmul 0 x = 0"  | 
|
1659  | 
by (simp add: frag_cmul_def)  | 
|
1660  | 
||
1661  | 
lemma frag_cmul_zero2 [simp]: "frag_cmul c 0 = 0"  | 
|
1662  | 
by (simp add: frag_cmul_def)  | 
|
1663  | 
||
1664  | 
lemma frag_cmul_one [simp]: "frag_cmul 1 x = x"  | 
|
1665  | 
by (auto simp: frag_cmul_def Poly_Mapping.poly_mapping.lookup_inverse)  | 
|
1666  | 
||
1667  | 
lemma frag_cmul_minus_one [simp]: "frag_cmul (-1) x = -x"  | 
|
1668  | 
by (simp add: frag_cmul_def uminus_poly_mapping_def poly_mapping_eqI)  | 
|
1669  | 
||
1670  | 
lemma frag_cmul_cmul [simp]: "frag_cmul c (frag_cmul d x) = frag_cmul (c*d) x"  | 
|
1671  | 
by (simp add: frag_cmul_def mult_ac)  | 
|
1672  | 
||
1673  | 
lemma lookup_frag_cmul [simp]: "poly_mapping.lookup (frag_cmul c x) i = c * poly_mapping.lookup x i"  | 
|
1674  | 
by (simp add: frag_cmul_def)  | 
|
1675  | 
||
1676  | 
lemma minus_frag_cmul [simp]: "- frag_cmul k x = frag_cmul (-k) x"  | 
|
1677  | 
by (simp add: poly_mapping_eqI)  | 
|
1678  | 
||
1679  | 
lemma keys_frag_of: "Poly_Mapping.keys(frag_of a) = {a}"
 | 
|
1680  | 
by simp  | 
|
1681  | 
||
1682  | 
lemma finite_cmul_nonzero: "finite {x. c * Poly_Mapping.lookup a x \<noteq> (0::int)}"
 | 
|
1683  | 
by simp  | 
|
1684  | 
||
1685  | 
lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \<subseteq> Poly_Mapping.keys a"  | 
|
1686  | 
using finite_cmul_nonzero [of c a]  | 
|
1687  | 
by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI)  | 
|
1688  | 
||
1689  | 
||
1690  | 
lemma keys_cmul_iff [iff]: "i \<in> Poly_Mapping.keys (frag_cmul c x) \<longleftrightarrow> i \<in> Poly_Mapping.keys x \<and> c \<noteq> 0"  | 
|
1691  | 
apply (auto simp: )  | 
|
1692  | 
apply (meson subsetD keys_cmul)  | 
|
1693  | 
by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff)  | 
|
1694  | 
||
1695  | 
lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a"  | 
|
1696  | 
by (metis (no_types, hide_lams) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym)  | 
|
1697  | 
||
1698  | 
lemma keys_diff:  | 
|
1699  | 
"Poly_Mapping.keys(a - b) \<subseteq> Poly_Mapping.keys a \<union> Poly_Mapping.keys b"  | 
|
1700  | 
by (auto simp add: in_keys_iff lookup_minus)  | 
|
1701  | 
||
1702  | 
lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \<longleftrightarrow> c = 0"
 | 
|
1703  | 
by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI)  | 
|
1704  | 
||
1705  | 
lemma frag_cmul_eq_0_iff [simp]: "frag_cmul k c = 0 \<longleftrightarrow> k=0 \<or> c=0"  | 
|
1706  | 
by auto (metis subsetI subset_antisym keys_cmul_iff keys_eq_empty)  | 
|
1707  | 
||
1708  | 
lemma frag_of_eq: "frag_of x = frag_of y \<longleftrightarrow> x = y"  | 
|
1709  | 
by (metis lookup_single_eq lookup_single_not_eq zero_neq_one)  | 
|
1710  | 
||
1711  | 
lemma frag_cmul_distrib: "frag_cmul (c+d) a = frag_cmul c a + frag_cmul d a"  | 
|
1712  | 
by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib)  | 
|
1713  | 
||
1714  | 
lemma frag_cmul_distrib2: "frag_cmul c (a+b) = frag_cmul c a + frag_cmul c b"  | 
|
1715  | 
proof -  | 
|
1716  | 
  have "finite {x. poly_mapping.lookup a x + poly_mapping.lookup b x \<noteq> 0}"
 | 
|
1717  | 
using keys_add [of a b]  | 
|
1718  | 
by (metis (no_types, lifting) finite_keys finite_subset keys.rep_eq lookup_add mem_Collect_eq subsetI)  | 
|
1719  | 
then show ?thesis  | 
|
1720  | 
by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib)  | 
|
1721  | 
qed  | 
|
1722  | 
||
1723  | 
lemma frag_cmul_diff_distrib: "frag_cmul (a - b) c = frag_cmul a c - frag_cmul b c"  | 
|
1724  | 
by (auto simp: left_diff_distrib lookup_minus poly_mapping_eqI)  | 
|
1725  | 
||
1726  | 
lemma frag_cmul_sum:  | 
|
1727  | 
"frag_cmul a (sum b I) = (\<Sum>i\<in>I. frag_cmul a (b i))"  | 
|
1728  | 
proof (induction rule: infinite_finite_induct)  | 
|
1729  | 
case (insert i I)  | 
|
1730  | 
then show ?case  | 
|
1731  | 
by (auto simp: algebra_simps frag_cmul_distrib2)  | 
|
1732  | 
qed auto  | 
|
1733  | 
||
1734  | 
lemma keys_sum: "Poly_Mapping.keys(sum b I) \<subseteq> (\<Union>i \<in>I. Poly_Mapping.keys(b i))"  | 
|
1735  | 
proof (induction I rule: infinite_finite_induct)  | 
|
1736  | 
case (insert i I)  | 
|
1737  | 
then show ?case  | 
|
1738  | 
using keys_add [of "b i" "sum b I"] by auto  | 
|
1739  | 
qed auto  | 
|
1740  | 
||
1741  | 
||
1742  | 
definition frag_extend :: "('b \<Rightarrow> 'a \<Rightarrow>\<^sub>0 int) \<Rightarrow> ('b \<Rightarrow>\<^sub>0 int) \<Rightarrow> 'a \<Rightarrow>\<^sub>0 int"
 | 
|
1743  | 
where "frag_extend b x \<equiv> (\<Sum>i \<in> Poly_Mapping.keys x. frag_cmul (Poly_Mapping.lookup x i) (b i))"  | 
|
1744  | 
||
1745  | 
lemma frag_extend_0 [simp]: "frag_extend b 0 = 0"  | 
|
1746  | 
by (simp add: frag_extend_def)  | 
|
1747  | 
||
1748  | 
lemma frag_extend_of [simp]: "frag_extend f (frag_of a) = f a"  | 
|
1749  | 
by (simp add: frag_extend_def)  | 
|
1750  | 
||
1751  | 
lemma frag_extend_cmul:  | 
|
1752  | 
"frag_extend f (frag_cmul c x) = frag_cmul c (frag_extend f x)"  | 
|
1753  | 
by (auto simp: frag_extend_def frag_cmul_sum intro: sum.mono_neutral_cong_left)  | 
|
1754  | 
||
1755  | 
lemma frag_extend_minus:  | 
|
1756  | 
"frag_extend f (- x) = - (frag_extend f x)"  | 
|
1757  | 
using frag_extend_cmul [of f "-1"] by simp  | 
|
1758  | 
||
1759  | 
lemma frag_extend_add:  | 
|
1760  | 
"frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)"  | 
|
1761  | 
proof -  | 
|
1762  | 
have *: "(\<Sum>i\<in>Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i))  | 
|
1763  | 
= (\<Sum>i\<in>Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))"  | 
|
1764  | 
"(\<Sum>i\<in>Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))  | 
|
1765  | 
= (\<Sum>i\<in>Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))"  | 
|
1766  | 
by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left)  | 
|
1767  | 
have "frag_extend f (a+b) = (\<Sum>i\<in>Poly_Mapping.keys (a + b).  | 
|
1768  | 
frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) "  | 
|
1769  | 
by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib)  | 
|
1770  | 
also have "... = (\<Sum>i \<in> Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i)  | 
|
1771  | 
+ frag_cmul (poly_mapping.lookup b i) (f i))"  | 
|
1772  | 
apply (rule sum.mono_neutral_cong_left)  | 
|
1773  | 
using keys_add [of a b]  | 
|
1774  | 
apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric])  | 
|
1775  | 
done  | 
|
1776  | 
also have "... = (frag_extend f a) + (frag_extend f b)"  | 
|
1777  | 
by (auto simp: * sum.distrib frag_extend_def)  | 
|
1778  | 
finally show ?thesis .  | 
|
1779  | 
qed  | 
|
1780  | 
||
1781  | 
lemma frag_extend_diff:  | 
|
1782  | 
"frag_extend f (a-b) = (frag_extend f a) - (frag_extend f b)"  | 
|
1783  | 
by (metis (no_types, hide_lams) add_uminus_conv_diff frag_extend_add frag_extend_minus)  | 
|
1784  | 
||
1785  | 
lemma frag_extend_sum:  | 
|
1786  | 
"finite I \<Longrightarrow> frag_extend f (\<Sum>i\<in>I. g i) = sum (frag_extend f o g) I"  | 
|
1787  | 
by (induction I rule: finite_induct) (simp_all add: frag_extend_add)  | 
|
1788  | 
||
1789  | 
lemma frag_extend_eq:  | 
|
1790  | 
"(\<And>f. f \<in> Poly_Mapping.keys c \<Longrightarrow> g f = h f) \<Longrightarrow> frag_extend g c = frag_extend h c"  | 
|
1791  | 
by (simp add: frag_extend_def)  | 
|
1792  | 
||
1793  | 
lemma frag_extend_eq_0:  | 
|
1794  | 
"(\<And>x. x \<in> Poly_Mapping.keys c \<Longrightarrow> f x = 0) \<Longrightarrow> frag_extend f c = 0"  | 
|
1795  | 
by (simp add: frag_extend_def)  | 
|
1796  | 
||
1797  | 
lemma keys_frag_extend: "Poly_Mapping.keys(frag_extend f c) \<subseteq> (\<Union>x \<in> Poly_Mapping.keys c. Poly_Mapping.keys(f x))"  | 
|
1798  | 
unfolding frag_extend_def  | 
|
1799  | 
by (smt SUP_mono' keys_cmul keys_sum order_trans)  | 
|
1800  | 
||
1801  | 
lemma frag_expansion: "a = frag_extend frag_of a"  | 
|
1802  | 
proof -  | 
|
1803  | 
have *: "finite I  | 
|
1804  | 
\<Longrightarrow> Poly_Mapping.lookup (\<Sum>i\<in>I. frag_cmul (Poly_Mapping.lookup a i) (frag_of i)) j =  | 
|
1805  | 
(if j \<in> I then Poly_Mapping.lookup a j else 0)" for I j  | 
|
1806  | 
by (induction I rule: finite_induct) (auto simp: lookup_single lookup_add)  | 
|
1807  | 
show ?thesis  | 
|
1808  | 
unfolding frag_extend_def  | 
|
1809  | 
by (rule poly_mapping_eqI) (fastforce simp add: in_keys_iff *)  | 
|
1810  | 
qed  | 
|
1811  | 
||
1812  | 
lemma frag_closure_minus_cmul:  | 
|
1813  | 
assumes "P 0" and P: "\<And>x y. \<lbrakk>P x; P y\<rbrakk> \<Longrightarrow> P(x - y)" "P c"  | 
|
1814  | 
shows "P(frag_cmul k c)"  | 
|
1815  | 
proof -  | 
|
1816  | 
have "P (frag_cmul (int n) c)" for n  | 
|
1817  | 
apply (induction n)  | 
|
1818  | 
apply (simp_all add: assms frag_cmul_distrib)  | 
|
1819  | 
by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P)  | 
|
1820  | 
then show ?thesis  | 
|
1821  | 
by (metis (no_types, hide_lams) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases)  | 
|
1822  | 
qed  | 
|
1823  | 
||
1824  | 
lemma frag_induction [consumes 1, case_names zero one diff]:  | 
|
1825  | 
assumes supp: "Poly_Mapping.keys c \<subseteq> S"  | 
|
1826  | 
and 0: "P 0" and sing: "\<And>x. x \<in> S \<Longrightarrow> P(frag_of x)"  | 
|
1827  | 
and diff: "\<And>a b. \<lbrakk>P a; P b\<rbrakk> \<Longrightarrow> P(a - b)"  | 
|
1828  | 
shows "P c"  | 
|
1829  | 
proof -  | 
|
1830  | 
have "P (\<Sum>i\<in>I. frag_cmul (poly_mapping.lookup c i) (frag_of i))"  | 
|
1831  | 
if "I \<subseteq> Poly_Mapping.keys c" for I  | 
|
1832  | 
using finite_subset [OF that finite_keys [of c]] that supp  | 
|
1833  | 
proof (induction I arbitrary: c rule: finite_induct)  | 
|
1834  | 
case empty  | 
|
1835  | 
then show ?case  | 
|
1836  | 
by (auto simp: 0)  | 
|
1837  | 
next  | 
|
1838  | 
case (insert i I c)  | 
|
1839  | 
have ab: "a+b = a - (0 - b)" for a b :: "'a \<Rightarrow>\<^sub>0 int"  | 
|
1840  | 
by simp  | 
|
1841  | 
have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))"  | 
|
1842  | 
by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff)  | 
|
1843  | 
show ?case  | 
|
1844  | 
apply (simp add: insert.hyps)  | 
|
1845  | 
apply (subst ab)  | 
|
1846  | 
using insert apply (blast intro: assms Pfrag)  | 
|
1847  | 
done  | 
|
1848  | 
qed  | 
|
1849  | 
then show ?thesis  | 
|
1850  | 
by (subst frag_expansion) (auto simp add: frag_extend_def)  | 
|
1851  | 
qed  | 
|
1852  | 
||
1853  | 
lemma frag_extend_compose:  | 
|
1854  | 
"frag_extend f (frag_extend (frag_of o g) c) = frag_extend (f o g) c"  | 
|
1855  | 
using subset_UNIV  | 
|
1856  | 
by (induction c rule: frag_induction) (auto simp: frag_extend_diff)  | 
|
1857  | 
||
1858  | 
lemma frag_split:  | 
|
1859  | 
fixes c :: "'a \<Rightarrow>\<^sub>0 int"  | 
|
1860  | 
assumes "Poly_Mapping.keys c \<subseteq> S \<union> T"  | 
|
1861  | 
obtains d e where "Poly_Mapping.keys d \<subseteq> S" "Poly_Mapping.keys e \<subseteq> T" "d + e = c"  | 
|
1862  | 
proof  | 
|
1863  | 
let ?d = "frag_extend (\<lambda>f. if f \<in> S then frag_of f else 0) c"  | 
|
1864  | 
let ?e = "frag_extend (\<lambda>f. if f \<in> S then 0 else frag_of f) c"  | 
|
1865  | 
show "Poly_Mapping.keys ?d \<subseteq> S" "Poly_Mapping.keys ?e \<subseteq> T"  | 
|
1866  | 
using assms by (auto intro!: order_trans [OF keys_frag_extend] split: if_split_asm)  | 
|
1867  | 
show "?d + ?e = c"  | 
|
1868  | 
using assms  | 
|
1869  | 
proof (induction c rule: frag_induction)  | 
|
1870  | 
case (diff a b)  | 
|
1871  | 
then show ?case  | 
|
1872  | 
by (metis (no_types, lifting) frag_extend_diff add_diff_eq diff_add_eq diff_add_eq_diff_diff_swap)  | 
|
1873  | 
qed auto  | 
|
1874  | 
qed  | 
|
1875  | 
||
1876  | 
hide_const (open) lookup single update keys range map map_key degree nth the_value items foldr mapp  | 
|
1877  | 
||
1878  | 
end  |