--- a/NEWS Thu Mar 11 19:05:46 2010 +0100
+++ b/NEWS Thu Mar 11 19:06:03 2010 +0100
@@ -83,6 +83,9 @@
*** HOL ***
+* Split off theory Big_Operators containing setsum, setprod, Inf_fin, Sup_fin,
+Min, Max from theory Finite_Set. INCOMPATIBILITY.
+
* Theory "Rational" renamed to "Rat", for consistency with "Nat", "Int" etc.
INCOMPATIBILITY.
@@ -102,7 +105,7 @@
INCOMPATIBILITY.
-* Class division ring also requires proof of fact divide_inverse. However instantiation
+* Class division_ring also requires proof of fact divide_inverse. However instantiation
of parameter divide has also been required previously. INCOMPATIBILITY.
* More consistent naming of type classes involving orderings (and lattices):
--- a/doc-src/Nitpick/nitpick.tex Thu Mar 11 19:05:46 2010 +0100
+++ b/doc-src/Nitpick/nitpick.tex Thu Mar 11 19:06:03 2010 +0100
@@ -137,8 +137,8 @@
suggesting several textual improvements.
% and Perry James for reporting a typo.
-\section{Overview}
-\label{overview}
+\section{First Steps}
+\label{first-steps}
This section introduces Nitpick by presenting small examples. If possible, you
should try out the examples on your workstation. Your theory file should start
@@ -472,7 +472,7 @@
\prew
\textbf{lemma} ``$\forall n.\; \textit{Suc}~n \mathbin{\not=} n \,\Longrightarrow\, P$'' \\
-\textbf{nitpick} [\textit{card~nat}~= 100, \textit{check\_potential}] \\[2\smallskipamount]
+\textbf{nitpick} [\textit{card~nat}~= 50, \textit{check\_potential}] \\[2\smallskipamount]
\slshape Warning: The conjecture either trivially holds for the given scopes or lies outside Nitpick's supported
fragment. Only potential counterexamples may be found. \\[2\smallskipamount]
Nitpick found a potential counterexample: \\[2\smallskipamount]
@@ -629,7 +629,7 @@
unlikely that one could be found for smaller cardinalities.
\subsection{Typedefs, Quotient Types, Records, Rationals, and Reals}
-\label{typedefs-records-rationals-and-reals}
+\label{typedefs-quotient-types-records-rationals-and-reals}
Nitpick generally treats types declared using \textbf{typedef} as datatypes
whose single constructor is the corresponding \textit{Abs\_\kern.1ex} function.
@@ -684,7 +684,26 @@
In the counterexample, $\Abs{(0,\, 0)}$ and $\Abs{(1,\, 0)}$ represent the
integers $0$ and $1$, respectively. Other representants would have been
-possible---e.g., $\Abs{(5,\, 5)}$ and $\Abs{(12,\, 11)}$.
+possible---e.g., $\Abs{(5,\, 5)}$ and $\Abs{(12,\, 11)}$. If we are going to
+use \textit{my\_int} extensively, it pays off to install a term postprocessor
+that converts the pair notation to the standard mathematical notation:
+
+\prew
+$\textbf{ML}~\,\{{*} \\
+\!\begin{aligned}[t]
+%& ({*}~\,\textit{Proof.context} \rightarrow \textit{string} \rightarrow (\textit{typ} \rightarrow \textit{term~list\/}) \rightarrow \textit{typ} \rightarrow \textit{term} \\[-2pt]
+%& \phantom{(*}~\,{\rightarrow}\;\textit{term}~\,{*}) \\[-2pt]
+& \textbf{fun}\,~\textit{my\_int\_postproc}~\_~\_~\_~T~(\textit{Const}~\_~\$~(\textit{Const}~\_~\$~\textit{t1}~\$~\textit{t2\/})) = {} \\[-2pt]
+& \phantom{fun}\,~\textit{HOLogic.mk\_number}~T~(\textit{snd}~(\textit{HOLogic.dest\_number~t1}) \\[-2pt]
+& \phantom{fun\,~\textit{HOLogic.mk\_number}~T~(}{-}~\textit{snd}~(\textit{HOLogic.dest\_number~t2\/})) \\[-2pt]
+& \phantom{fun}\!{\mid}\,~\textit{my\_int\_postproc}~\_~\_~\_~\_~t = t \\[-2pt]
+{*}\}\end{aligned}$ \\[2\smallskipamount]
+$\textbf{setup}~\,\{{*} \\
+\!\begin{aligned}[t]
+& \textit{Nitpick.register\_term\_postprocessor}~\!\begin{aligned}[t]
+ & @\{\textrm{typ}~\textit{my\_int}\}~\textit{my\_int\_postproc}\end{aligned} \\[-2pt]
+{*}\}\end{aligned}$
+\postw
Records are also handled as datatypes with a single constructor:
@@ -771,25 +790,25 @@
\prew
\textbf{lemma} ``$\exists n.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
-\textbf{nitpick}~[\textit{card nat}~= 100, \textit{unary\_ints}, \textit{verbose}] \\[2\smallskipamount]
+\textbf{nitpick}~[\textit{card nat}~= 50, \textit{unary\_ints}, \textit{verbose}] \\[2\smallskipamount]
\slshape The inductive predicate ``\textit{even}'' was proved well-founded.
Nitpick can compute it efficiently. \\[2\smallskipamount]
Trying 1 scope: \\
-\hbox{}\qquad \textit{card nat}~= 100. \\[2\smallskipamount]
-Nitpick found a potential counterexample for \textit{card nat}~= 100: \\[2\smallskipamount]
+\hbox{}\qquad \textit{card nat}~= 50. \\[2\smallskipamount]
+Nitpick found a potential counterexample for \textit{card nat}~= 50: \\[2\smallskipamount]
\hbox{}\qquad Empty assignment \\[2\smallskipamount]
Nitpick could not find a better counterexample. \\[2\smallskipamount]
Total time: 2274 ms.
\postw
No genuine counterexample is possible because Nitpick cannot rule out the
-existence of a natural number $n \ge 100$ such that both $\textit{even}~n$ and
+existence of a natural number $n \ge 50$ such that both $\textit{even}~n$ and
$\textit{even}~(\textit{Suc}~n)$ are true. To help Nitpick, we can bound the
existential quantifier:
\prew
-\textbf{lemma} ``$\exists n \mathbin{\le} 99.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
-\textbf{nitpick}~[\textit{card nat}~= 100, \textit{unary\_ints}] \\[2\smallskipamount]
+\textbf{lemma} ``$\exists n \mathbin{\le} 49.\; \textit{even}~n \mathrel{\land} \textit{even}~(\textit{Suc}~n)$'' \\
+\textbf{nitpick}~[\textit{card nat}~= 50, \textit{unary\_ints}] \\[2\smallskipamount]
\slshape Nitpick found a counterexample: \\[2\smallskipamount]
\hbox{}\qquad Empty assignment
\postw
@@ -1211,26 +1230,26 @@
Trying 8 scopes: \\
\hbox{}\qquad \textit{card} $'a$~= 1, \textit{card} $'b$~= 1,
\textit{card} \textit{nat}~= 1, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 1, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 1, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 1. \\
+\textit{list\/}''~= 1, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 1, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 1. \\
\hbox{}\qquad \textit{card} $'a$~= 2, \textit{card} $'b$~= 2,
\textit{card} \textit{nat}~= 2, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 2, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 2, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 2. \\
+\textit{list\/}''~= 2, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 2, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 2. \\
\hbox{}\qquad $\qquad\vdots$ \\[.5\smallskipamount]
\hbox{}\qquad \textit{card} $'a$~= 8, \textit{card} $'b$~= 8,
\textit{card} \textit{nat}~= 8, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 8, \\
-\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list}''~= 8, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 8.
+\textit{list\/}''~= 8, \\
+\hbox{}\qquad\quad \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 8, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 8.
\\[2\smallskipamount]
Nitpick found a counterexample for
\textit{card} $'a$~= 5, \textit{card} $'b$~= 5,
\textit{card} \textit{nat}~= 5, \textit{card} ``$('a \times {'}b)$
-\textit{list}''~= 5, \textit{card} ``\kern1pt$'a$ \textit{list}''~= 5, and
-\textit{card} ``\kern1pt$'b$ \textit{list}''~= 5:
+\textit{list\/}''~= 5, \textit{card} ``\kern1pt$'a$ \textit{list\/}''~= 5, and
+\textit{card} ``\kern1pt$'b$ \textit{list\/}''~= 5:
\\[2\smallskipamount]
\hbox{}\qquad Free variables: \nopagebreak \\
\hbox{}\qquad\qquad $\textit{xs} = [a_1, a_2]$ \\
@@ -1851,7 +1870,7 @@
The number of options can be overwhelming at first glance. Do not let that worry
you: Nitpick's defaults have been chosen so that it almost always does the right
thing, and the most important options have been covered in context in
-\S\ref{overview}.
+\S\ref{first-steps}.
The descriptions below refer to the following syntactic quantities:
@@ -2622,7 +2641,7 @@
then Nitpick assumes that the definition was made using an inductive package and
based on the introduction rules marked with \textit{nitpick\_\allowbreak
-ind\_\allowbreak intros} tries to determine whether the definition is
+\allowbreak intros} tries to determine whether the definition is
well-founded.
\end{enum}
\end{enum}
@@ -2664,7 +2683,8 @@
Nitpick provides a rich Standard ML interface used mainly for internal purposes
and debugging. Among the most interesting functions exported by Nitpick are
those that let you invoke the tool programmatically and those that let you
-register and unregister custom coinductive datatypes.
+register and unregister custom coinductive datatypes as well as term
+postprocessors.
\subsection{Invocation of Nitpick}
\label{invocation-of-nitpick}
@@ -2695,7 +2715,7 @@
put into a \textit{params} record. Here is an example:
\prew
-$\textbf{val}\,~\textit{params} = \textit{Nitpick\_Isar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
+$\textbf{val}\,~\textit{params} = \textit{Nitpick\_Isar.default\_params}~\textit{thy}~[(\textrm{``}\textrm{timeout\/}\textrm{''},\, \textrm{``}\textrm{none}\textrm{''})]$ \\
$\textbf{val}\,~(\textit{outcome},\, \textit{state}') = \textit{Nitpick.pick\_nits\_in\_subgoal}~\begin{aligned}[t]
& \textit{state}~\textit{params}~\textit{false} \\[-2pt]
& \textit{subgoal}\end{aligned}$
@@ -2726,7 +2746,7 @@
\prew
$\textbf{setup}~\,\{{*}\,~\!\begin{aligned}[t]
& \textit{Nitpick.register\_codatatype} \\[-2pt]
-& \qquad @\{\antiq{typ}~``\kern1pt'a~\textit{llist}\textrm{''}\}~@\{\antiq{const\_name}~ \textit{llist\_case}\} \\[-2pt] %% TYPESETTING
+& \qquad @\{\antiq{typ}~``\kern1pt'a~\textit{llist\/}\textrm{''}\}~@\{\antiq{const\_name}~ \textit{llist\_case}\} \\[-2pt] %% TYPESETTING
& \qquad (\textit{map}~\textit{dest\_Const}~[@\{\antiq{term}~\textit{LNil}\},\, @\{\antiq{term}~\textit{LCons}\}])\,\ {*}\}\end{aligned}$
\postw
@@ -2740,7 +2760,7 @@
\prew
$\textbf{setup}~\,\{{*}\,~\textit{Nitpick.unregister\_codatatype}~@\{\antiq{typ}~``
-\kern1pt'a~\textit{list}\textrm{''}\}\ \,{*}\}$
+\kern1pt'a~\textit{list\/}\textrm{''}\}\ \,{*}\}$
\postw
Inductive datatypes can be registered as coinductive datatypes, given
@@ -2748,6 +2768,26 @@
the use of the inductive constructors---Nitpick will generate an error if they
are needed.
+\subsection{Registration of Term Postprocessors}
+\label{registration-of-term-postprocessors}
+
+It is possible to change the output of any term that Nitpick considers a
+datatype by registering a term postprocessor. The interface for registering and
+unregistering postprocessors consists of the following pair of functions defined
+in the \textit{Nitpick} structure:
+
+\prew
+$\textbf{type}\,~\textit{term\_postprocessor}\,~{=} {}$ \\
+$\hbox{}\quad\textit{Proof.context} \rightarrow \textit{string} \rightarrow (\textit{typ} \rightarrow \textit{term~list\/}) \rightarrow \textit{typ} \rightarrow \textit{term} \rightarrow \textit{term}$ \\
+$\textbf{val}\,~\textit{register\_term\_postprocessors} : {}$ \\
+$\hbox{}\quad\textit{typ} \rightarrow \textit{term\_postprocessor} \rightarrow \textit{theory} \rightarrow \textit{theory}$ \\
+$\textbf{val}\,~\textit{unregister\_term\_postprocessors} :\,
+\textit{typ} \rightarrow \textit{theory} \rightarrow \textit{theory}$
+\postw
+
+\S\ref{typedefs-quotient-types-records-rationals-and-reals} and
+\texttt{src/HOL/Library/Multiset.thy} illustrate this feature in context.
+
\section{Known Bugs and Limitations}
\label{known-bugs-and-limitations}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Big_Operators.thy Thu Mar 11 19:06:03 2010 +0100
@@ -0,0 +1,1711 @@
+(* Title: HOL/Big_Operators.thy
+ Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
+ with contributions by Jeremy Avigad
+*)
+
+header {* Big operators and finite (non-empty) sets *}
+
+theory Big_Operators
+imports Plain
+begin
+
+subsection {* Generalized summation over a set *}
+
+interpretation comm_monoid_add: comm_monoid_mult "op +" "0::'a::comm_monoid_add"
+ proof qed (auto intro: add_assoc add_commute)
+
+definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
+where "setsum f A == if finite A then fold_image (op +) f 0 A else 0"
+
+abbreviation
+ Setsum ("\<Sum>_" [1000] 999) where
+ "\<Sum>A == setsum (%x. x) A"
+
+text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
+written @{text"\<Sum>x\<in>A. e"}. *}
+
+syntax
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10)
+syntax (xsymbols)
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
+syntax (HTML output)
+ "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
+
+translations -- {* Beware of argument permutation! *}
+ "SUM i:A. b" == "CONST setsum (%i. b) A"
+ "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
+
+text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Sum>x|P. e"}. *}
+
+syntax
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
+syntax (xsymbols)
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
+syntax (HTML output)
+ "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
+
+translations
+ "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
+ "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
+
+print_translation {*
+let
+ fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
+ if x <> y then raise Match
+ else
+ let
+ val x' = Syntax.mark_bound x;
+ val t' = subst_bound (x', t);
+ val P' = subst_bound (x', P);
+ in Syntax.const @{syntax_const "_qsetsum"} $ Syntax.mark_bound x $ P' $ t' end
+ | setsum_tr' _ = raise Match;
+in [(@{const_syntax setsum}, setsum_tr')] end
+*}
+
+
+lemma setsum_empty [simp]: "setsum f {} = 0"
+by (simp add: setsum_def)
+
+lemma setsum_insert [simp]:
+ "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
+by (simp add: setsum_def)
+
+lemma setsum_infinite [simp]: "~ finite A ==> setsum f A = 0"
+by (simp add: setsum_def)
+
+lemma setsum_reindex:
+ "inj_on f B ==> setsum h (f ` B) = setsum (h \<circ> f) B"
+by(auto simp add: setsum_def comm_monoid_add.fold_image_reindex dest!:finite_imageD)
+
+lemma setsum_reindex_id:
+ "inj_on f B ==> setsum f B = setsum id (f ` B)"
+by (auto simp add: setsum_reindex)
+
+lemma setsum_reindex_nonzero:
+ assumes fS: "finite S"
+ and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
+ shows "setsum h (f ` S) = setsum (h o f) S"
+using nz
+proof(induct rule: finite_induct[OF fS])
+ case 1 thus ?case by simp
+next
+ case (2 x F)
+ {assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
+ then obtain y where y: "y \<in> F" "f x = f y" by auto
+ from "2.hyps" y have xy: "x \<noteq> y" by auto
+
+ from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
+ have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
+ also have "\<dots> = setsum (h o f) (insert x F)"
+ unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+ using h0
+ apply simp
+ apply (rule "2.hyps"(3))
+ apply (rule_tac y="y" in "2.prems")
+ apply simp_all
+ done
+ finally have ?case .}
+ moreover
+ {assume fxF: "f x \<notin> f ` F"
+ have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"
+ using fxF "2.hyps" by simp
+ also have "\<dots> = setsum (h o f) (insert x F)"
+ unfolding setsum_insert[OF `finite F` `x\<notin>F`]
+ apply simp
+ apply (rule cong[OF refl[of "op + (h (f x))"]])
+ apply (rule "2.hyps"(3))
+ apply (rule_tac y="y" in "2.prems")
+ apply simp_all
+ done
+ finally have ?case .}
+ ultimately show ?case by blast
+qed
+
+lemma setsum_cong:
+ "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
+by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong)
+
+lemma strong_setsum_cong[cong]:
+ "A = B ==> (!!x. x:B =simp=> f x = g x)
+ ==> setsum (%x. f x) A = setsum (%x. g x) B"
+by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong)
+
+lemma setsum_cong2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setsum f A = setsum g A"
+by (rule setsum_cong[OF refl], auto)
+
+lemma setsum_reindex_cong:
+ "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]
+ ==> setsum h B = setsum g A"
+by (simp add: setsum_reindex cong: setsum_cong)
+
+
+lemma setsum_0[simp]: "setsum (%i. 0) A = 0"
+apply (clarsimp simp: setsum_def)
+apply (erule finite_induct, auto)
+done
+
+lemma setsum_0': "ALL a:A. f a = 0 ==> setsum f A = 0"
+by(simp add:setsum_cong)
+
+lemma setsum_Un_Int: "finite A ==> finite B ==>
+ setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
+ -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
+by(simp add: setsum_def comm_monoid_add.fold_image_Un_Int [symmetric])
+
+lemma setsum_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
+by (subst setsum_Un_Int [symmetric], auto)
+
+lemma setsum_mono_zero_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 0"
+ shows "setsum f S = setsum f T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_right:
+ "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. f i = 0 \<Longrightarrow> setsum f T = setsum f S"
+by(blast intro!: setsum_mono_zero_left[symmetric])
+
+lemma setsum_mono_zero_cong_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. g i = 0"
+ and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+ shows "setsum f S = setsum g T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
+qed
+
+lemma setsum_mono_zero_cong_right:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 0"
+ and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
+ shows "setsum f T = setsum g S"
+using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto
+
+lemma setsum_delta:
+ assumes fS: "finite S"
+ shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else 0)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = 0" by simp
+ hence ?thesis using a by simp}
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
+ using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a by simp}
+ ultimately show ?thesis by blast
+qed
+lemma setsum_delta':
+ assumes fS: "finite S" shows
+ "setsum (\<lambda>k. if a = k then b k else 0) S =
+ (if a\<in> S then b a else 0)"
+ using setsum_delta[OF fS, of a b, symmetric]
+ by (auto intro: setsum_cong)
+
+lemma setsum_restrict_set:
+ assumes fA: "finite A"
+ shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
+proof-
+ from fA have fab: "finite (A \<inter> B)" by auto
+ have aba: "A \<inter> B \<subseteq> A" by blast
+ let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
+ from setsum_mono_zero_left[OF fA aba, of ?g]
+ show ?thesis by simp
+qed
+
+lemma setsum_cases:
+ assumes fA: "finite A"
+ shows "setsum (\<lambda>x. if P x then f x else g x) A =
+ setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
+proof-
+ have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}"
+ "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
+ by blast+
+ from fA
+ have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
+ let ?g = "\<lambda>x. if P x then f x else g x"
+ from setsum_Un_disjoint[OF f a(2), of ?g] a(1)
+ show ?thesis by simp
+qed
+
+
+(*But we can't get rid of finite I. If infinite, although the rhs is 0,
+ the lhs need not be, since UNION I A could still be finite.*)
+lemma setsum_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
+ setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
+by(simp add: setsum_def comm_monoid_add.fold_image_UN_disjoint cong: setsum_cong)
+
+text{*No need to assume that @{term C} is finite. If infinite, the rhs is
+directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*}
+lemma setsum_Union_disjoint:
+ "[| (ALL A:C. finite A);
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
+ ==> setsum f (Union C) = setsum (setsum f) C"
+apply (cases "finite C")
+ prefer 2 apply (force dest: finite_UnionD simp add: setsum_def)
+ apply (frule setsum_UN_disjoint [of C id f])
+ apply (unfold Union_def id_def, assumption+)
+done
+
+(*But we can't get rid of finite A. If infinite, although the lhs is 0,
+ the rhs need not be, since SIGMA A B could still be finite.*)
+lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
+ (\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"
+by(simp add:setsum_def comm_monoid_add.fold_image_Sigma split_def cong:setsum_cong)
+
+text{*Here we can eliminate the finiteness assumptions, by cases.*}
+lemma setsum_cartesian_product:
+ "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
+apply (cases "finite A")
+ apply (cases "finite B")
+ apply (simp add: setsum_Sigma)
+ apply (cases "A={}", simp)
+ apply (simp)
+apply (auto simp add: setsum_def
+ dest: finite_cartesian_productD1 finite_cartesian_productD2)
+done
+
+lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
+by(simp add:setsum_def comm_monoid_add.fold_image_distrib)
+
+
+subsubsection {* Properties in more restricted classes of structures *}
+
+lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
+apply (case_tac "finite A")
+ prefer 2 apply (simp add: setsum_def)
+apply (erule rev_mp)
+apply (erule finite_induct, auto)
+done
+
+lemma setsum_eq_0_iff [simp]:
+ "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
+by (induct set: finite) auto
+
+lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>
+ (setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"
+apply(erule finite_induct)
+apply (auto simp add:add_is_1)
+done
+
+lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
+
+lemma setsum_Un_nat: "finite A ==> finite B ==>
+ (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
+ -- {* For the natural numbers, we have subtraction. *}
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
+
+lemma setsum_Un: "finite A ==> finite B ==>
+ (setsum f (A Un B) :: 'a :: ab_group_add) =
+ setsum f A + setsum f B - setsum f (A Int B)"
+by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
+
+lemma (in comm_monoid_mult) fold_image_1: "finite S \<Longrightarrow> (\<forall>x\<in>S. f x = 1) \<Longrightarrow> fold_image op * f 1 S = 1"
+ apply (induct set: finite)
+ apply simp by auto
+
+lemma (in comm_monoid_mult) fold_image_Un_one:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
+ shows "fold_image (op *) f 1 (S \<union> T) = fold_image (op *) f 1 S * fold_image (op *) f 1 T"
+proof-
+ have "fold_image op * f 1 (S \<inter> T) = 1"
+ apply (rule fold_image_1)
+ using fS fT I0 by auto
+ with fold_image_Un_Int[OF fS fT] show ?thesis by simp
+qed
+
+lemma setsum_eq_general_reverses:
+ assumes fS: "finite S" and fT: "finite T"
+ and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
+ and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"
+ shows "setsum f S = setsum g T"
+ apply (simp add: setsum_def fS fT)
+ apply (rule comm_monoid_add.fold_image_eq_general_inverses[OF fS])
+ apply (erule kh)
+ apply (erule hk)
+ done
+
+
+
+lemma setsum_Un_zero:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 0"
+ shows "setsum f (S \<union> T) = setsum f S + setsum f T"
+ using fS fT
+ apply (simp add: setsum_def)
+ apply (rule comm_monoid_add.fold_image_Un_one)
+ using I0 by auto
+
+
+lemma setsum_UNION_zero:
+ assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
+ and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"
+ shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
+ using fSS f0
+proof(induct rule: finite_induct[OF fS])
+ case 1 thus ?case by simp
+next
+ case (2 T F)
+ then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"
+ and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
+ from fTF have fUF: "finite (\<Union>F)" by auto
+ from "2.prems" TF fTF
+ show ?case
+ by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f])
+qed
+
+
+lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
+ (if a:A then setsum f A - f a else setsum f A)"
+apply (case_tac "finite A")
+ prefer 2 apply (simp add: setsum_def)
+apply (erule finite_induct)
+ apply (auto simp add: insert_Diff_if)
+apply (drule_tac a = a in mk_disjoint_insert, auto)
+done
+
+lemma setsum_diff1: "finite A \<Longrightarrow>
+ (setsum f (A - {a}) :: ('a::ab_group_add)) =
+ (if a:A then setsum f A - f a else setsum f A)"
+by (erule finite_induct) (auto simp add: insert_Diff_if)
+
+lemma setsum_diff1'[rule_format]:
+ "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
+apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
+apply (auto simp add: insert_Diff_if add_ac)
+done
+
+lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
+ shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
+unfolding setsum_diff1'[OF assms] by auto
+
+(* By Jeremy Siek: *)
+
+lemma setsum_diff_nat:
+assumes "finite B" and "B \<subseteq> A"
+shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"
+using assms
+proof induct
+ show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
+next
+ fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"
+ and xFinA: "insert x F \<subseteq> A"
+ and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"
+ from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp
+ from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
+ by (simp add: setsum_diff1_nat)
+ from xFinA have "F \<subseteq> A" by simp
+ with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
+ with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
+ by simp
+ from xnotinF have "A - insert x F = (A - F) - {x}" by auto
+ with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
+ by simp
+ from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp
+ with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
+ by simp
+ thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp
+qed
+
+lemma setsum_diff:
+ assumes le: "finite A" "B \<subseteq> A"
+ shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
+proof -
+ from le have finiteB: "finite B" using finite_subset by auto
+ show ?thesis using finiteB le
+ proof induct
+ case empty
+ thus ?case by auto
+ next
+ case (insert x F)
+ thus ?case using le finiteB
+ by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
+ qed
+qed
+
+lemma setsum_mono:
+ assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
+ shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
+proof (cases "finite K")
+ case True
+ thus ?thesis using le
+ proof induct
+ case empty
+ thus ?case by simp
+ next
+ case insert
+ thus ?case using add_mono by fastsimp
+ qed
+next
+ case False
+ thus ?thesis
+ by (simp add: setsum_def)
+qed
+
+lemma setsum_strict_mono:
+ fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
+ assumes "finite A" "A \<noteq> {}"
+ and "!!x. x:A \<Longrightarrow> f x < g x"
+ shows "setsum f A < setsum g A"
+ using prems
+proof (induct rule: finite_ne_induct)
+ case singleton thus ?case by simp
+next
+ case insert thus ?case by (auto simp: add_strict_mono)
+qed
+
+lemma setsum_negf:
+ "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"
+proof (cases "finite A")
+ case True thus ?thesis by (induct set: finite) auto
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_subtractf:
+ "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
+ setsum f A - setsum g A"
+proof (cases "finite A")
+ case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_nonneg:
+ assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
+ shows "0 \<le> setsum f A"
+proof (cases "finite A")
+ case True thus ?thesis using nn
+ proof induct
+ case empty then show ?case by simp
+ next
+ case (insert x F)
+ then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
+ with insert show ?case by simp
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_nonpos:
+ assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
+ shows "setsum f A \<le> 0"
+proof (cases "finite A")
+ case True thus ?thesis using np
+ proof induct
+ case empty then show ?case by simp
+ next
+ case (insert x F)
+ then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
+ with insert show ?case by simp
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_mono2:
+fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_semigroup_add_imp_le,comm_monoid_add}"
+assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
+shows "setsum f A \<le> setsum f B"
+proof -
+ have "setsum f A \<le> setsum f A + setsum f (B-A)"
+ by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
+ also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
+ by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
+ also have "A \<union> (B-A) = B" using sub by blast
+ finally show ?thesis .
+qed
+
+lemma setsum_mono3: "finite B ==> A <= B ==>
+ ALL x: B - A.
+ 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
+ setsum f A <= setsum f B"
+ apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")
+ apply (erule ssubst)
+ apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")
+ apply simp
+ apply (rule add_left_mono)
+ apply (erule setsum_nonneg)
+ apply (subst setsum_Un_disjoint [THEN sym])
+ apply (erule finite_subset, assumption)
+ apply (rule finite_subset)
+ prefer 2
+ apply assumption
+ apply (auto simp add: sup_absorb2)
+done
+
+lemma setsum_right_distrib:
+ fixes f :: "'a => ('b::semiring_0)"
+ shows "r * setsum f A = setsum (%n. r * f n) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: right_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_left_distrib:
+ "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"
+proof (cases "finite A")
+ case True
+ then show ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: left_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_divide_distrib:
+ "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"
+proof (cases "finite A")
+ case True
+ then show ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (simp add: add_divide_distrib)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_abs[iff]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A)
+ thus ?case by (auto intro: abs_triangle_ineq order_trans)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma setsum_abs_ge_zero[iff]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "0 \<le> setsum (%i. abs(f i)) A"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg)
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+lemma abs_setsum_abs[simp]:
+ fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
+ shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"
+proof (cases "finite A")
+ case True
+ thus ?thesis
+ proof induct
+ case empty thus ?case by simp
+ next
+ case (insert a A)
+ hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp
+ also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp
+ also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"
+ by (simp del: abs_of_nonneg)
+ also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp
+ finally show ?case .
+ qed
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+
+lemma setsum_Plus:
+ fixes A :: "'a set" and B :: "'b set"
+ assumes fin: "finite A" "finite B"
+ shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
+proof -
+ have "A <+> B = Inl ` A \<union> Inr ` B" by auto
+ moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
+ by(auto intro: finite_imageI)
+ moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
+ moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
+ ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex)
+qed
+
+
+text {* Commuting outer and inner summation *}
+
+lemma swap_inj_on:
+ "inj_on (%(i, j). (j, i)) (A \<times> B)"
+ by (unfold inj_on_def) fast
+
+lemma swap_product:
+ "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
+ by (simp add: split_def image_def) blast
+
+lemma setsum_commute:
+ "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
+proof (simp add: setsum_cartesian_product)
+ have "(\<Sum>(x,y) \<in> A <*> B. f x y) =
+ (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"
+ (is "?s = _")
+ apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on)
+ apply (simp add: split_def)
+ done
+ also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"
+ (is "_ = ?t")
+ apply (simp add: swap_product)
+ done
+ finally show "?s = ?t" .
+qed
+
+lemma setsum_product:
+ fixes f :: "'a => ('b::semiring_0)"
+ shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
+ by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
+
+lemma setsum_mult_setsum_if_inj:
+fixes f :: "'a => ('b::semiring_0)"
+shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
+ setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
+by(auto simp: setsum_product setsum_cartesian_product
+ intro!: setsum_reindex_cong[symmetric])
+
+lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"
+apply (cases "finite A")
+apply (erule finite_induct)
+apply (auto simp add: algebra_simps)
+done
+
+lemma setsum_bounded:
+ assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
+ shows "setsum f A \<le> of_nat(card A) * K"
+proof (cases "finite A")
+ case True
+ thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp
+next
+ case False thus ?thesis by (simp add: setsum_def)
+qed
+
+
+subsubsection {* Cardinality as special case of @{const setsum} *}
+
+lemma card_eq_setsum:
+ "card A = setsum (\<lambda>x. 1) A"
+ by (simp only: card_def setsum_def)
+
+lemma card_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {})
+ ==> card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
+apply (simp add: card_eq_setsum del: setsum_constant)
+apply (subgoal_tac
+ "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I")
+apply (simp add: setsum_UN_disjoint del: setsum_constant)
+apply (simp cong: setsum_cong)
+done
+
+lemma card_Union_disjoint:
+ "finite C ==> (ALL A:C. finite A) ==>
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
+ ==> card (Union C) = setsum card C"
+apply (frule card_UN_disjoint [of C id])
+apply (unfold Union_def id_def, assumption+)
+done
+
+text{*The image of a finite set can be expressed using @{term fold_image}.*}
+lemma image_eq_fold_image:
+ "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A"
+proof (induct rule: finite_induct)
+ case empty then show ?case by simp
+next
+ interpret ab_semigroup_mult "op Un"
+ proof qed auto
+ case insert
+ then show ?case by simp
+qed
+
+subsubsection {* Cardinality of products *}
+
+lemma card_SigmaI [simp]:
+ "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
+ \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
+by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant)
+
+(*
+lemma SigmaI_insert: "y \<notin> A ==>
+ (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
+ by auto
+*)
+
+lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"
+ by (cases "finite A \<and> finite B")
+ (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2)
+
+lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)"
+by (simp add: card_cartesian_product)
+
+
+subsection {* Generalized product over a set *}
+
+definition setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult"
+where "setprod f A == if finite A then fold_image (op *) f 1 A else 1"
+
+abbreviation
+ Setprod ("\<Prod>_" [1000] 999) where
+ "\<Prod>A == setprod (%x. x) A"
+
+syntax
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10)
+syntax (xsymbols)
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
+syntax (HTML output)
+ "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
+
+translations -- {* Beware of argument permutation! *}
+ "PROD i:A. b" == "CONST setprod (%i. b) A"
+ "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
+
+text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
+ @{text"\<Prod>x|P. e"}. *}
+
+syntax
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
+syntax (xsymbols)
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
+syntax (HTML output)
+ "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
+
+translations
+ "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
+ "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
+
+
+lemma setprod_empty [simp]: "setprod f {} = 1"
+by (auto simp add: setprod_def)
+
+lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
+ setprod f (insert a A) = f a * setprod f A"
+by (simp add: setprod_def)
+
+lemma setprod_infinite [simp]: "~ finite A ==> setprod f A = 1"
+by (simp add: setprod_def)
+
+lemma setprod_reindex:
+ "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
+by(auto simp: setprod_def fold_image_reindex dest!:finite_imageD)
+
+lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)"
+by (auto simp add: setprod_reindex)
+
+lemma setprod_cong:
+ "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
+by(fastsimp simp: setprod_def intro: fold_image_cong)
+
+lemma strong_setprod_cong[cong]:
+ "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
+by(fastsimp simp: simp_implies_def setprod_def intro: fold_image_cong)
+
+lemma setprod_reindex_cong: "inj_on f A ==>
+ B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
+by (frule setprod_reindex, simp)
+
+lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"
+ and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
+ shows "setprod h B = setprod g A"
+proof-
+ have "setprod h B = setprod (h o f) A"
+ by (simp add: B setprod_reindex[OF i, of h])
+ then show ?thesis apply simp
+ apply (rule setprod_cong)
+ apply simp
+ by (simp add: eq)
+qed
+
+lemma setprod_Un_one:
+ assumes fS: "finite S" and fT: "finite T"
+ and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
+ shows "setprod f (S \<union> T) = setprod f S * setprod f T"
+ using fS fT
+ apply (simp add: setprod_def)
+ apply (rule fold_image_Un_one)
+ using I0 by auto
+
+
+lemma setprod_1: "setprod (%i. 1) A = 1"
+apply (case_tac "finite A")
+apply (erule finite_induct, auto simp add: mult_ac)
+done
+
+lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1"
+apply (subgoal_tac "setprod f F = setprod (%x. 1) F")
+apply (erule ssubst, rule setprod_1)
+apply (rule setprod_cong, auto)
+done
+
+lemma setprod_Un_Int: "finite A ==> finite B
+ ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
+by(simp add: setprod_def fold_image_Un_Int[symmetric])
+
+lemma setprod_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
+by (subst setprod_Un_Int [symmetric], auto)
+
+lemma setprod_mono_one_left:
+ assumes fT: "finite T" and ST: "S \<subseteq> T"
+ and z: "\<forall>i \<in> T - S. f i = 1"
+ shows "setprod f S = setprod f T"
+proof-
+ have eq: "T = S \<union> (T - S)" using ST by blast
+ have d: "S \<inter> (T - S) = {}" using ST by blast
+ from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
+ show ?thesis
+ by (simp add: setprod_Un_disjoint[OF f d, unfolded eq[symmetric]] setprod_1'[OF z])
+qed
+
+lemmas setprod_mono_one_right = setprod_mono_one_left [THEN sym]
+
+lemma setprod_delta:
+ assumes fS: "finite S"
+ shows "setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else 1)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = 1" by simp
+ hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) }
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto
+ have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+ using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)}
+ ultimately show ?thesis by blast
+qed
+
+lemma setprod_delta':
+ assumes fS: "finite S" shows
+ "setprod (\<lambda>k. if a = k then b k else 1) S =
+ (if a\<in> S then b a else 1)"
+ using setprod_delta[OF fS, of a b, symmetric]
+ by (auto intro: setprod_cong)
+
+
+lemma setprod_UN_disjoint:
+ "finite I ==> (ALL i:I. finite (A i)) ==>
+ (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
+ setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
+by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong)
+
+lemma setprod_Union_disjoint:
+ "[| (ALL A:C. finite A);
+ (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
+ ==> setprod f (Union C) = setprod (setprod f) C"
+apply (cases "finite C")
+ prefer 2 apply (force dest: finite_UnionD simp add: setprod_def)
+ apply (frule setprod_UN_disjoint [of C id f])
+ apply (unfold Union_def id_def, assumption+)
+done
+
+lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
+ (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
+ (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
+by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong)
+
+text{*Here we can eliminate the finiteness assumptions, by cases.*}
+lemma setprod_cartesian_product:
+ "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
+apply (cases "finite A")
+ apply (cases "finite B")
+ apply (simp add: setprod_Sigma)
+ apply (cases "A={}", simp)
+ apply (simp add: setprod_1)
+apply (auto simp add: setprod_def
+ dest: finite_cartesian_productD1 finite_cartesian_productD2)
+done
+
+lemma setprod_timesf:
+ "setprod (%x. f x * g x) A = (setprod f A * setprod g A)"
+by(simp add:setprod_def fold_image_distrib)
+
+
+subsubsection {* Properties in more restricted classes of structures *}
+
+lemma setprod_eq_1_iff [simp]:
+ "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"
+by (induct set: finite) auto
+
+lemma setprod_zero:
+ "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"
+apply (induct set: finite, force, clarsimp)
+apply (erule disjE, auto)
+done
+
+lemma setprod_nonneg [rule_format]:
+ "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
+by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)
+
+lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)
+ --> 0 < setprod f A"
+by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)
+
+lemma setprod_zero_iff[simp]: "finite A ==>
+ (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
+ (EX x: A. f x = 0)"
+by (erule finite_induct, auto simp:no_zero_divisors)
+
+lemma setprod_pos_nat:
+ "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
+using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
+
+lemma setprod_pos_nat_iff[simp]:
+ "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"
+using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
+
+lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
+ (setprod f (A Un B) :: 'a ::{field})
+ = setprod f A * setprod f B / setprod f (A Int B)"
+by (subst setprod_Un_Int [symmetric], auto)
+
+lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
+ (setprod f (A - {a}) :: 'a :: {field}) =
+ (if a:A then setprod f A / f a else setprod f A)"
+by (erule finite_induct) (auto simp add: insert_Diff_if)
+
+lemma setprod_inversef:
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
+by (erule finite_induct) auto
+
+lemma setprod_dividef:
+ fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+ shows "finite A
+ ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
+apply (subgoal_tac
+ "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
+apply (erule ssubst)
+apply (subst divide_inverse)
+apply (subst setprod_timesf)
+apply (subst setprod_inversef, assumption+, rule refl)
+apply (rule setprod_cong, rule refl)
+apply (subst divide_inverse, auto)
+done
+
+lemma setprod_dvd_setprod [rule_format]:
+ "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply (auto simp add: dvd_def)
+ apply (rule_tac x = "k * ka" in exI)
+ apply (simp add: algebra_simps)
+done
+
+lemma setprod_dvd_setprod_subset:
+ "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
+ apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
+ apply (unfold dvd_def, blast)
+ apply (subst setprod_Un_disjoint [symmetric])
+ apply (auto elim: finite_subset intro: setprod_cong)
+done
+
+lemma setprod_dvd_setprod_subset2:
+ "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>
+ setprod f A dvd setprod g B"
+ apply (rule dvd_trans)
+ apply (rule setprod_dvd_setprod, erule (1) bspec)
+ apply (erule (1) setprod_dvd_setprod_subset)
+done
+
+lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>
+ (f i ::'a::comm_semiring_1) dvd setprod f A"
+by (induct set: finite) (auto intro: dvd_mult)
+
+lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>
+ (d::'a::comm_semiring_1) dvd (SUM x : A. f x)"
+ apply (cases "finite A")
+ apply (induct set: finite)
+ apply auto
+done
+
+lemma setprod_mono:
+ fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"
+ assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"
+ shows "setprod f A \<le> setprod g A"
+proof (cases "finite A")
+ case True
+ hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]
+ proof (induct A rule: finite_subset_induct)
+ case (insert a F)
+ thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
+ unfolding setprod_insert[OF insert(1,3)]
+ using assms[rule_format,OF insert(2)] insert
+ by (auto intro: mult_mono mult_nonneg_nonneg)
+ qed auto
+ thus ?thesis by simp
+qed auto
+
+lemma abs_setprod:
+ fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
+ shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"
+proof (cases "finite A")
+ case True thus ?thesis
+ by induct (auto simp add: field_simps abs_mult)
+qed auto
+
+lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
+apply (erule finite_induct)
+apply auto
+done
+
+lemma setprod_gen_delta:
+ assumes fS: "finite S"
+ shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult}) * c^ (card S - 1) else c^ card S)"
+proof-
+ let ?f = "(\<lambda>k. if k=a then b k else c)"
+ {assume a: "a \<notin> S"
+ hence "\<forall> k\<in> S. ?f k = c" by simp
+ hence ?thesis using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) }
+ moreover
+ {assume a: "a \<in> S"
+ let ?A = "S - {a}"
+ let ?B = "{a}"
+ have eq: "S = ?A \<union> ?B" using a by blast
+ have dj: "?A \<inter> ?B = {}" by simp
+ from fS have fAB: "finite ?A" "finite ?B" by auto
+ have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
+ apply (rule setprod_cong) by auto
+ have cA: "card ?A = card S - 1" using fS a by auto
+ have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto
+ have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
+ using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
+ by simp
+ then have ?thesis using a cA
+ by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
+ ultimately show ?thesis by blast
+qed
+
+
+subsubsection {* Fold1 in lattices with @{const inf} and @{const sup} *}
+
+text{*
+ As an application of @{text fold1} we define infimum
+ and supremum in (not necessarily complete!) lattices
+ over (non-empty) sets by means of @{text fold1}.
+*}
+
+context semilattice_inf
+begin
+
+lemma below_fold1_iff:
+ assumes "finite A" "A \<noteq> {}"
+ shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
+qed
+
+lemma fold1_belowI:
+ assumes "finite A"
+ and "a \<in> A"
+ shows "fold1 inf A \<le> a"
+proof -
+ from assms have "A \<noteq> {}" by auto
+ from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
+ proof (induct rule: finite_ne_induct)
+ case singleton thus ?case by simp
+ next
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ case (insert x F)
+ from insert(5) have "a = x \<or> a \<in> F" by simp
+ thus ?case
+ proof
+ assume "a = x" thus ?thesis using insert
+ by (simp add: mult_ac)
+ next
+ assume "a \<in> F"
+ hence bel: "fold1 inf F \<le> a" by (rule insert)
+ have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"
+ using insert by (simp add: mult_ac)
+ also have "inf (fold1 inf F) a = fold1 inf F"
+ using bel by (auto intro: antisym)
+ also have "inf x \<dots> = fold1 inf (insert x F)"
+ using insert by (simp add: mult_ac)
+ finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .
+ moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp
+ ultimately show ?thesis by simp
+ qed
+ qed
+qed
+
+end
+
+context lattice
+begin
+
+definition
+ Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900)
+where
+ "Inf_fin = fold1 inf"
+
+definition
+ Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900)
+where
+ "Sup_fin = fold1 sup"
+
+lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A"
+apply(unfold Sup_fin_def Inf_fin_def)
+apply(subgoal_tac "EX a. a:A")
+prefer 2 apply blast
+apply(erule exE)
+apply(rule order_trans)
+apply(erule (1) fold1_belowI)
+apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice])
+done
+
+lemma sup_Inf_absorb [simp]:
+ "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a"
+apply(subst sup_commute)
+apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI)
+done
+
+lemma inf_Sup_absorb [simp]:
+ "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"
+by (simp add: Sup_fin_def inf_absorb1
+ semilattice_inf.fold1_belowI [OF dual_semilattice])
+
+end
+
+context distrib_lattice
+begin
+
+lemma sup_Inf1_distrib:
+ assumes "finite A"
+ and "A \<noteq> {}"
+ shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ from assms show ?thesis
+ by (simp add: Inf_fin_def image_def
+ hom_fold1_commute [where h="sup x", OF sup_inf_distrib1])
+ (rule arg_cong [where f="fold1 inf"], blast)
+qed
+
+lemma sup_Inf2_distrib:
+ assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
+ shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}"
+using A proof (induct rule: finite_ne_induct)
+ case singleton thus ?case
+ by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
+next
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ case (insert x A)
+ have finB: "finite {sup x b |b. b \<in> B}"
+ by(rule finite_surj[where f = "sup x", OF B(1)], auto)
+ have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
+ proof -
+ have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
+ by blast
+ thus ?thesis by(simp add: insert(1) B(1))
+ qed
+ have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
+ have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)"
+ using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def])
+ also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2)
+ also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})"
+ using insert by(simp add:sup_Inf1_distrib[OF B])
+ also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
+ (is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M")
+ using B insert
+ by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne])
+ also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
+ by blast
+ finally show ?case .
+qed
+
+lemma inf_Sup1_distrib:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
+proof -
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ from assms show ?thesis
+ by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
+ (rule arg_cong [where f="fold1 sup"], blast)
+qed
+
+lemma inf_Sup2_distrib:
+ assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
+ shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}"
+using A proof (induct rule: finite_ne_induct)
+ case singleton thus ?case
+ by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def])
+next
+ case (insert x A)
+ have finB: "finite {inf x b |b. b \<in> B}"
+ by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)
+ have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
+ proof -
+ have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
+ by blast
+ thus ?thesis by(simp add: insert(1) B(1))
+ qed
+ have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
+ using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
+ also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2)
+ also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})"
+ using insert by(simp add:inf_Sup1_distrib[OF B])
+ also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
+ (is "_ = \<Squnion>\<^bsub>fin\<^esub>?M")
+ using B insert
+ by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne])
+ also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
+ by blast
+ finally show ?case .
+qed
+
+end
+
+context complete_lattice
+begin
+
+lemma Inf_fin_Inf:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
+proof -
+ interpret ab_semigroup_idem_mult inf
+ by (rule ab_semigroup_idem_mult_inf)
+ from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+ moreover with `finite A` have "finite B" by simp
+ ultimately show ?thesis
+ by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
+ (simp add: Inf_fold_inf)
+qed
+
+lemma Sup_fin_Sup:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
+proof -
+ interpret ab_semigroup_idem_mult sup
+ by (rule ab_semigroup_idem_mult_sup)
+ from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
+ moreover with `finite A` have "finite B" by simp
+ ultimately show ?thesis
+ by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])
+ (simp add: Sup_fold_sup)
+qed
+
+end
+
+
+subsubsection {* Fold1 in linear orders with @{const min} and @{const max} *}
+
+text{*
+ As an application of @{text fold1} we define minimum
+ and maximum in (not necessarily complete!) linear orders
+ over (non-empty) sets by means of @{text fold1}.
+*}
+
+context linorder
+begin
+
+lemma ab_semigroup_idem_mult_min:
+ "ab_semigroup_idem_mult min"
+ proof qed (auto simp add: min_def)
+
+lemma ab_semigroup_idem_mult_max:
+ "ab_semigroup_idem_mult max"
+ proof qed (auto simp add: max_def)
+
+lemma max_lattice:
+ "semilattice_inf (op \<ge>) (op >) max"
+ by (fact min_max.dual_semilattice)
+
+lemma dual_max:
+ "ord.max (op \<ge>) = min"
+ by (auto simp add: ord.max_def_raw min_def expand_fun_eq)
+
+lemma dual_min:
+ "ord.min (op \<ge>) = max"
+ by (auto simp add: ord.min_def_raw max_def expand_fun_eq)
+
+lemma strict_below_fold1_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert)
+qed
+
+lemma fold1_below_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert min_le_iff_disj)
+qed
+
+lemma fold1_strict_below_iff:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (induct rule: finite_ne_induct)
+ (simp_all add: fold1_insert min_less_iff_disj)
+qed
+
+lemma fold1_antimono:
+ assumes "A \<noteq> {}" and "A \<subseteq> B" and "finite B"
+ shows "fold1 min B \<le> fold1 min A"
+proof cases
+ assume "A = B" thus ?thesis by simp
+next
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ assume "A \<noteq> B"
+ have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
+ have "fold1 min B = fold1 min (A \<union> (B-A))" by(subst B)(rule refl)
+ also have "\<dots> = min (fold1 min A) (fold1 min (B-A))"
+ proof -
+ have "finite A" by(rule finite_subset[OF `A \<subseteq> B` `finite B`])
+ moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *)
+ moreover have "(B-A) \<noteq> {}" using prems by blast
+ moreover have "A Int (B-A) = {}" using prems by blast
+ ultimately show ?thesis using `A \<noteq> {}` by (rule_tac fold1_Un)
+ qed
+ also have "\<dots> \<le> fold1 min A" by (simp add: min_le_iff_disj)
+ finally show ?thesis .
+qed
+
+definition
+ Min :: "'a set \<Rightarrow> 'a"
+where
+ "Min = fold1 min"
+
+definition
+ Max :: "'a set \<Rightarrow> 'a"
+where
+ "Max = fold1 max"
+
+lemmas Min_singleton [simp] = fold1_singleton_def [OF Min_def]
+lemmas Max_singleton [simp] = fold1_singleton_def [OF Max_def]
+
+lemma Min_insert [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min (insert x A) = min x (Min A)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
+qed
+
+lemma Max_insert [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max (insert x A) = max x (Max A)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
+qed
+
+lemma Min_in [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A \<in> A"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
+qed
+
+lemma Max_in [simp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A \<in> A"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
+qed
+
+lemma Min_Un:
+ assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
+ shows "Min (A \<union> B) = min (Min A) (Min B)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (simp add: Min_def fold1_Un2)
+qed
+
+lemma Max_Un:
+ assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
+ shows "Max (A \<union> B) = max (Max A) (Max B)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis
+ by (simp add: Max_def fold1_Un2)
+qed
+
+lemma hom_Min_commute:
+ assumes "\<And>x y. h (min x y) = min (h x) (h y)"
+ and "finite N" and "N \<noteq> {}"
+ shows "h (Min N) = Min (h ` N)"
+proof -
+ interpret ab_semigroup_idem_mult min
+ by (rule ab_semigroup_idem_mult_min)
+ from assms show ?thesis
+ by (simp add: Min_def hom_fold1_commute)
+qed
+
+lemma hom_Max_commute:
+ assumes "\<And>x y. h (max x y) = max (h x) (h y)"
+ and "finite N" and "N \<noteq> {}"
+ shows "h (Max N) = Max (h ` N)"
+proof -
+ interpret ab_semigroup_idem_mult max
+ by (rule ab_semigroup_idem_mult_max)
+ from assms show ?thesis
+ by (simp add: Max_def hom_fold1_commute [of h])
+qed
+
+lemma Min_le [simp]:
+ assumes "finite A" and "x \<in> A"
+ shows "Min A \<le> x"
+ using assms by (simp add: Min_def min_max.fold1_belowI)
+
+lemma Max_ge [simp]:
+ assumes "finite A" and "x \<in> A"
+ shows "x \<le> Max A"
+proof -
+ interpret semilattice_inf "op \<ge>" "op >" max
+ by (rule max_lattice)
+ from assms show ?thesis by (simp add: Max_def fold1_belowI)
+qed
+
+lemma Min_ge_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
+ using assms by (simp add: Min_def min_max.below_fold1_iff)
+
+lemma Max_le_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
+proof -
+ interpret semilattice_inf "op \<ge>" "op >" max
+ by (rule max_lattice)
+ from assms show ?thesis by (simp add: Max_def below_fold1_iff)
+qed
+
+lemma Min_gr_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
+ using assms by (simp add: Min_def strict_below_fold1_iff)
+
+lemma Max_less_iff [simp, noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.strict_below_fold1_iff [folded dual.dual_max])
+qed
+
+lemma Min_le_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
+ using assms by (simp add: Min_def fold1_below_iff)
+
+lemma Max_ge_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_below_iff [folded dual.dual_max])
+qed
+
+lemma Min_less_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
+ using assms by (simp add: Min_def fold1_strict_below_iff)
+
+lemma Max_gr_iff [noatp]:
+ assumes "finite A" and "A \<noteq> {}"
+ shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_strict_below_iff [folded dual.dual_max])
+qed
+
+lemma Min_eqI:
+ assumes "finite A"
+ assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"
+ and "x \<in> A"
+ shows "Min A = x"
+proof (rule antisym)
+ from `x \<in> A` have "A \<noteq> {}" by auto
+ with assms show "Min A \<ge> x" by simp
+next
+ from assms show "x \<ge> Min A" by simp
+qed
+
+lemma Max_eqI:
+ assumes "finite A"
+ assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
+ and "x \<in> A"
+ shows "Max A = x"
+proof (rule antisym)
+ from `x \<in> A` have "A \<noteq> {}" by auto
+ with assms show "Max A \<le> x" by simp
+next
+ from assms show "x \<le> Max A" by simp
+qed
+
+lemma Min_antimono:
+ assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
+ shows "Min N \<le> Min M"
+ using assms by (simp add: Min_def fold1_antimono)
+
+lemma Max_mono:
+ assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
+ shows "Max M \<le> Max N"
+proof -
+ interpret dual: linorder "op \<ge>" "op >"
+ by (rule dual_linorder)
+ from assms show ?thesis
+ by (simp add: Max_def dual.fold1_antimono [folded dual.dual_max])
+qed
+
+lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:
+ "finite A \<Longrightarrow> P {} \<Longrightarrow>
+ (!!b A. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A))
+ \<Longrightarrow> P A"
+proof (induct rule: finite_psubset_induct)
+ fix A :: "'a set"
+ assume IH: "!! B. finite B \<Longrightarrow> B < A \<Longrightarrow> P {} \<Longrightarrow>
+ (!!b A. finite A \<Longrightarrow> (\<forall>a\<in>A. a<b) \<Longrightarrow> P A \<Longrightarrow> P (insert b A))
+ \<Longrightarrow> P B"
+ and "finite A" and "P {}"
+ and step: "!!b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)"
+ show "P A"
+ proof (cases "A = {}")
+ assume "A = {}" thus "P A" using `P {}` by simp
+ next
+ let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B"
+ assume "A \<noteq> {}"
+ with `finite A` have "Max A : A" by auto
+ hence A: "?A = A" using insert_Diff_single insert_absorb by auto
+ moreover have "finite ?B" using `finite A` by simp
+ ultimately have "P ?B" using `P {}` step IH[of ?B] by blast
+ moreover have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp
+ ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp
+ qed
+qed
+
+lemma finite_linorder_min_induct[consumes 1, case_names empty insert]:
+ "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
+by(rule linorder.finite_linorder_max_induct[OF dual_linorder])
+
+end
+
+context linordered_ab_semigroup_add
+begin
+
+lemma add_Min_commute:
+ fixes k
+ assumes "finite N" and "N \<noteq> {}"
+ shows "k + Min N = Min {k + m | m. m \<in> N}"
+proof -
+ have "\<And>x y. k + min x y = min (k + x) (k + y)"
+ by (simp add: min_def not_le)
+ (blast intro: antisym less_imp_le add_left_mono)
+ with assms show ?thesis
+ using hom_Min_commute [of "plus k" N]
+ by simp (blast intro: arg_cong [where f = Min])
+qed
+
+lemma add_Max_commute:
+ fixes k
+ assumes "finite N" and "N \<noteq> {}"
+ shows "k + Max N = Max {k + m | m. m \<in> N}"
+proof -
+ have "\<And>x y. k + max x y = max (k + x) (k + y)"
+ by (simp add: max_def not_le)
+ (blast intro: antisym less_imp_le add_left_mono)
+ with assms show ?thesis
+ using hom_Max_commute [of "plus k" N]
+ by simp (blast intro: arg_cong [where f = Max])
+qed
+
+end
+
+context linordered_ab_group_add
+begin
+
+lemma minus_Max_eq_Min [simp]:
+ "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Max S) = Min (uminus ` S)"
+ by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)
+
+lemma minus_Min_eq_Max [simp]:
+ "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Min S) = Max (uminus ` S)"
+ by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)
+
+end
+
+end
--- a/src/HOL/Equiv_Relations.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Equiv_Relations.thy Thu Mar 11 19:06:03 2010 +0100
@@ -5,7 +5,7 @@
header {* Equivalence Relations in Higher-Order Set Theory *}
theory Equiv_Relations
-imports Finite_Set Relation Plain
+imports Big_Operators Relation Plain
begin
subsection {* Equivalence relations *}
--- a/src/HOL/Finite_Set.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Finite_Set.thy Thu Mar 11 19:06:03 2010 +0100
@@ -6,7 +6,7 @@
header {* Finite sets *}
theory Finite_Set
-imports Power Product_Type Sum_Type
+imports Power Option
begin
subsection {* Definition and basic properties *}
@@ -527,17 +527,24 @@
lemma UNIV_unit [noatp]:
"UNIV = {()}" by auto
-instance unit :: finite
- by default (simp add: UNIV_unit)
+instance unit :: finite proof
+qed (simp add: UNIV_unit)
lemma UNIV_bool [noatp]:
"UNIV = {False, True}" by auto
-instance bool :: finite
- by default (simp add: UNIV_bool)
+instance bool :: finite proof
+qed (simp add: UNIV_bool)
+
+instance * :: (finite, finite) finite proof
+qed (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
-instance * :: (finite, finite) finite
- by default (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite)
+lemma finite_option_UNIV [simp]:
+ "finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"
+ by (auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some)
+
+instance option :: (finite) finite proof
+qed (simp add: UNIV_option_conv)
lemma inj_graph: "inj (%f. {(x, y). y = f x})"
by (rule inj_onI, auto simp add: expand_set_eq expand_fun_eq)
@@ -556,8 +563,8 @@
qed
qed
-instance "+" :: (finite, finite) finite
- by default (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
+instance "+" :: (finite, finite) finite proof
+qed (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite)
subsection {* A fold functional for finite sets *}
@@ -1053,1470 +1060,6 @@
end
-subsection {* Generalized summation over a set *}
-
-interpretation comm_monoid_add: comm_monoid_mult "op +" "0::'a::comm_monoid_add"
- proof qed (auto intro: add_assoc add_commute)
-
-definition setsum :: "('a => 'b) => 'a set => 'b::comm_monoid_add"
-where "setsum f A == if finite A then fold_image (op +) f 0 A else 0"
-
-abbreviation
- Setsum ("\<Sum>_" [1000] 999) where
- "\<Sum>A == setsum (%x. x) A"
-
-text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
-written @{text"\<Sum>x\<in>A. e"}. *}
-
-syntax
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10)
-syntax (xsymbols)
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
-syntax (HTML output)
- "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
-
-translations -- {* Beware of argument permutation! *}
- "SUM i:A. b" == "CONST setsum (%i. b) A"
- "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
-
-text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Sum>x|P. e"}. *}
-
-syntax
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
-syntax (xsymbols)
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
-syntax (HTML output)
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
-
-translations
- "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
- "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
-
-print_translation {*
-let
- fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
- if x <> y then raise Match
- else
- let
- val x' = Syntax.mark_bound x;
- val t' = subst_bound (x', t);
- val P' = subst_bound (x', P);
- in Syntax.const @{syntax_const "_qsetsum"} $ Syntax.mark_bound x $ P' $ t' end
- | setsum_tr' _ = raise Match;
-in [(@{const_syntax setsum}, setsum_tr')] end
-*}
-
-
-lemma setsum_empty [simp]: "setsum f {} = 0"
-by (simp add: setsum_def)
-
-lemma setsum_insert [simp]:
- "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
-by (simp add: setsum_def)
-
-lemma setsum_infinite [simp]: "~ finite A ==> setsum f A = 0"
-by (simp add: setsum_def)
-
-lemma setsum_reindex:
- "inj_on f B ==> setsum h (f ` B) = setsum (h \<circ> f) B"
-by(auto simp add: setsum_def comm_monoid_add.fold_image_reindex dest!:finite_imageD)
-
-lemma setsum_reindex_id:
- "inj_on f B ==> setsum f B = setsum id (f ` B)"
-by (auto simp add: setsum_reindex)
-
-lemma setsum_reindex_nonzero:
- assumes fS: "finite S"
- and nz: "\<And> x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
- shows "setsum h (f ` S) = setsum (h o f) S"
-using nz
-proof(induct rule: finite_induct[OF fS])
- case 1 thus ?case by simp
-next
- case (2 x F)
- {assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
- then obtain y where y: "y \<in> F" "f x = f y" by auto
- from "2.hyps" y have xy: "x \<noteq> y" by auto
-
- from "2.prems"[of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
- have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
- also have "\<dots> = setsum (h o f) (insert x F)"
- unfolding setsum_insert[OF `finite F` `x\<notin>F`]
- using h0
- apply simp
- apply (rule "2.hyps"(3))
- apply (rule_tac y="y" in "2.prems")
- apply simp_all
- done
- finally have ?case .}
- moreover
- {assume fxF: "f x \<notin> f ` F"
- have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"
- using fxF "2.hyps" by simp
- also have "\<dots> = setsum (h o f) (insert x F)"
- unfolding setsum_insert[OF `finite F` `x\<notin>F`]
- apply simp
- apply (rule cong[OF refl[of "op + (h (f x))"]])
- apply (rule "2.hyps"(3))
- apply (rule_tac y="y" in "2.prems")
- apply simp_all
- done
- finally have ?case .}
- ultimately show ?case by blast
-qed
-
-lemma setsum_cong:
- "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
-by(fastsimp simp: setsum_def intro: comm_monoid_add.fold_image_cong)
-
-lemma strong_setsum_cong[cong]:
- "A = B ==> (!!x. x:B =simp=> f x = g x)
- ==> setsum (%x. f x) A = setsum (%x. g x) B"
-by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong)
-
-lemma setsum_cong2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setsum f A = setsum g A"
-by (rule setsum_cong[OF refl], auto)
-
-lemma setsum_reindex_cong:
- "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]
- ==> setsum h B = setsum g A"
-by (simp add: setsum_reindex cong: setsum_cong)
-
-
-lemma setsum_0[simp]: "setsum (%i. 0) A = 0"
-apply (clarsimp simp: setsum_def)
-apply (erule finite_induct, auto)
-done
-
-lemma setsum_0': "ALL a:A. f a = 0 ==> setsum f A = 0"
-by(simp add:setsum_cong)
-
-lemma setsum_Un_Int: "finite A ==> finite B ==>
- setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
- -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
-by(simp add: setsum_def comm_monoid_add.fold_image_Un_Int [symmetric])
-
-lemma setsum_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
-by (subst setsum_Un_Int [symmetric], auto)
-
-lemma setsum_mono_zero_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 0"
- shows "setsum f S = setsum f T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
-qed
-
-lemma setsum_mono_zero_right:
- "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. f i = 0 \<Longrightarrow> setsum f T = setsum f S"
-by(blast intro!: setsum_mono_zero_left[symmetric])
-
-lemma setsum_mono_zero_cong_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. g i = 0"
- and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
- shows "setsum f S = setsum g T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- using fg by (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] setsum_0'[OF z])
-qed
-
-lemma setsum_mono_zero_cong_right:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 0"
- and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
- shows "setsum f T = setsum g S"
-using setsum_mono_zero_cong_left[OF fT ST z] fg[symmetric] by auto
-
-lemma setsum_delta:
- assumes fS: "finite S"
- shows "setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else 0)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = 0" by simp
- hence ?thesis using a by simp}
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have "setsum ?f S = setsum ?f ?A + setsum ?f ?B"
- using setsum_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a by simp}
- ultimately show ?thesis by blast
-qed
-lemma setsum_delta':
- assumes fS: "finite S" shows
- "setsum (\<lambda>k. if a = k then b k else 0) S =
- (if a\<in> S then b a else 0)"
- using setsum_delta[OF fS, of a b, symmetric]
- by (auto intro: setsum_cong)
-
-lemma setsum_restrict_set:
- assumes fA: "finite A"
- shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
-proof-
- from fA have fab: "finite (A \<inter> B)" by auto
- have aba: "A \<inter> B \<subseteq> A" by blast
- let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
- from setsum_mono_zero_left[OF fA aba, of ?g]
- show ?thesis by simp
-qed
-
-lemma setsum_cases:
- assumes fA: "finite A"
- shows "setsum (\<lambda>x. if P x then f x else g x) A =
- setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
-proof-
- have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}"
- "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
- by blast+
- from fA
- have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
- let ?g = "\<lambda>x. if P x then f x else g x"
- from setsum_Un_disjoint[OF f a(2), of ?g] a(1)
- show ?thesis by simp
-qed
-
-
-(*But we can't get rid of finite I. If infinite, although the rhs is 0,
- the lhs need not be, since UNION I A could still be finite.*)
-lemma setsum_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
- setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
-by(simp add: setsum_def comm_monoid_add.fold_image_UN_disjoint cong: setsum_cong)
-
-text{*No need to assume that @{term C} is finite. If infinite, the rhs is
-directly 0, and @{term "Union C"} is also infinite, hence the lhs is also 0.*}
-lemma setsum_Union_disjoint:
- "[| (ALL A:C. finite A);
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
- ==> setsum f (Union C) = setsum (setsum f) C"
-apply (cases "finite C")
- prefer 2 apply (force dest: finite_UnionD simp add: setsum_def)
- apply (frule setsum_UN_disjoint [of C id f])
- apply (unfold Union_def id_def, assumption+)
-done
-
-(*But we can't get rid of finite A. If infinite, although the lhs is 0,
- the rhs need not be, since SIGMA A B could still be finite.*)
-lemma setsum_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
- (\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"
-by(simp add:setsum_def comm_monoid_add.fold_image_Sigma split_def cong:setsum_cong)
-
-text{*Here we can eliminate the finiteness assumptions, by cases.*}
-lemma setsum_cartesian_product:
- "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
-apply (cases "finite A")
- apply (cases "finite B")
- apply (simp add: setsum_Sigma)
- apply (cases "A={}", simp)
- apply (simp)
-apply (auto simp add: setsum_def
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
-by(simp add:setsum_def comm_monoid_add.fold_image_distrib)
-
-
-subsubsection {* Properties in more restricted classes of structures *}
-
-lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
-apply (case_tac "finite A")
- prefer 2 apply (simp add: setsum_def)
-apply (erule rev_mp)
-apply (erule finite_induct, auto)
-done
-
-lemma setsum_eq_0_iff [simp]:
- "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
-by (induct set: finite) auto
-
-lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>
- (setsum f A = Suc 0) = (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"
-apply(erule finite_induct)
-apply (auto simp add:add_is_1)
-done
-
-lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
-
-lemma setsum_Un_nat: "finite A ==> finite B ==>
- (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
- -- {* For the natural numbers, we have subtraction. *}
-by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
-
-lemma setsum_Un: "finite A ==> finite B ==>
- (setsum f (A Un B) :: 'a :: ab_group_add) =
- setsum f A + setsum f B - setsum f (A Int B)"
-by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
-
-lemma (in comm_monoid_mult) fold_image_1: "finite S \<Longrightarrow> (\<forall>x\<in>S. f x = 1) \<Longrightarrow> fold_image op * f 1 S = 1"
- apply (induct set: finite)
- apply simp by auto
-
-lemma (in comm_monoid_mult) fold_image_Un_one:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
- shows "fold_image (op *) f 1 (S \<union> T) = fold_image (op *) f 1 S * fold_image (op *) f 1 T"
-proof-
- have "fold_image op * f 1 (S \<inter> T) = 1"
- apply (rule fold_image_1)
- using fS fT I0 by auto
- with fold_image_Un_Int[OF fS fT] show ?thesis by simp
-qed
-
-lemma setsum_eq_general_reverses:
- assumes fS: "finite S" and fT: "finite T"
- and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
- and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"
- shows "setsum f S = setsum g T"
- apply (simp add: setsum_def fS fT)
- apply (rule comm_monoid_add.fold_image_eq_general_inverses[OF fS])
- apply (erule kh)
- apply (erule hk)
- done
-
-
-
-lemma setsum_Un_zero:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 0"
- shows "setsum f (S \<union> T) = setsum f S + setsum f T"
- using fS fT
- apply (simp add: setsum_def)
- apply (rule comm_monoid_add.fold_image_Un_one)
- using I0 by auto
-
-
-lemma setsum_UNION_zero:
- assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
- and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"
- shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
- using fSS f0
-proof(induct rule: finite_induct[OF fS])
- case 1 thus ?case by simp
-next
- case (2 T F)
- then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"
- and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
- from fTF have fUF: "finite (\<Union>F)" by auto
- from "2.prems" TF fTF
- show ?case
- by (auto simp add: H[symmetric] intro: setsum_Un_zero[OF fTF(1) fUF, of f])
-qed
-
-
-lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
- (if a:A then setsum f A - f a else setsum f A)"
-apply (case_tac "finite A")
- prefer 2 apply (simp add: setsum_def)
-apply (erule finite_induct)
- apply (auto simp add: insert_Diff_if)
-apply (drule_tac a = a in mk_disjoint_insert, auto)
-done
-
-lemma setsum_diff1: "finite A \<Longrightarrow>
- (setsum f (A - {a}) :: ('a::ab_group_add)) =
- (if a:A then setsum f A - f a else setsum f A)"
-by (erule finite_induct) (auto simp add: insert_Diff_if)
-
-lemma setsum_diff1'[rule_format]:
- "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
-apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
-apply (auto simp add: insert_Diff_if add_ac)
-done
-
-lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
- shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
-unfolding setsum_diff1'[OF assms] by auto
-
-(* By Jeremy Siek: *)
-
-lemma setsum_diff_nat:
-assumes "finite B" and "B \<subseteq> A"
-shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"
-using assms
-proof induct
- show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
-next
- fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"
- and xFinA: "insert x F \<subseteq> A"
- and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"
- from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp
- from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
- by (simp add: setsum_diff1_nat)
- from xFinA have "F \<subseteq> A" by simp
- with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
- with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
- by simp
- from xnotinF have "A - insert x F = (A - F) - {x}" by auto
- with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
- by simp
- from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp
- with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
- by simp
- thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp
-qed
-
-lemma setsum_diff:
- assumes le: "finite A" "B \<subseteq> A"
- shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
-proof -
- from le have finiteB: "finite B" using finite_subset by auto
- show ?thesis using finiteB le
- proof induct
- case empty
- thus ?case by auto
- next
- case (insert x F)
- thus ?case using le finiteB
- by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
- qed
-qed
-
-lemma setsum_mono:
- assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
- shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
-proof (cases "finite K")
- case True
- thus ?thesis using le
- proof induct
- case empty
- thus ?case by simp
- next
- case insert
- thus ?case using add_mono by fastsimp
- qed
-next
- case False
- thus ?thesis
- by (simp add: setsum_def)
-qed
-
-lemma setsum_strict_mono:
- fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
- assumes "finite A" "A \<noteq> {}"
- and "!!x. x:A \<Longrightarrow> f x < g x"
- shows "setsum f A < setsum g A"
- using prems
-proof (induct rule: finite_ne_induct)
- case singleton thus ?case by simp
-next
- case insert thus ?case by (auto simp: add_strict_mono)
-qed
-
-lemma setsum_negf:
- "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"
-proof (cases "finite A")
- case True thus ?thesis by (induct set: finite) auto
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_subtractf:
- "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
- setsum f A - setsum g A"
-proof (cases "finite A")
- case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_nonneg:
- assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
- shows "0 \<le> setsum f A"
-proof (cases "finite A")
- case True thus ?thesis using nn
- proof induct
- case empty then show ?case by simp
- next
- case (insert x F)
- then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
- with insert show ?case by simp
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_nonpos:
- assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
- shows "setsum f A \<le> 0"
-proof (cases "finite A")
- case True thus ?thesis using np
- proof induct
- case empty then show ?case by simp
- next
- case (insert x F)
- then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
- with insert show ?case by simp
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_mono2:
-fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_semigroup_add_imp_le,comm_monoid_add}"
-assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
-shows "setsum f A \<le> setsum f B"
-proof -
- have "setsum f A \<le> setsum f A + setsum f (B-A)"
- by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
- also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
- by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
- also have "A \<union> (B-A) = B" using sub by blast
- finally show ?thesis .
-qed
-
-lemma setsum_mono3: "finite B ==> A <= B ==>
- ALL x: B - A.
- 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
- setsum f A <= setsum f B"
- apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")
- apply (erule ssubst)
- apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")
- apply simp
- apply (rule add_left_mono)
- apply (erule setsum_nonneg)
- apply (subst setsum_Un_disjoint [THEN sym])
- apply (erule finite_subset, assumption)
- apply (rule finite_subset)
- prefer 2
- apply assumption
- apply (auto simp add: sup_absorb2)
-done
-
-lemma setsum_right_distrib:
- fixes f :: "'a => ('b::semiring_0)"
- shows "r * setsum f A = setsum (%n. r * f n) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: right_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_left_distrib:
- "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"
-proof (cases "finite A")
- case True
- then show ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: left_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_divide_distrib:
- "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"
-proof (cases "finite A")
- case True
- then show ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (simp add: add_divide_distrib)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_abs[iff]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A)
- thus ?case by (auto intro: abs_triangle_ineq order_trans)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma setsum_abs_ge_zero[iff]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "0 \<le> setsum (%i. abs(f i)) A"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert x A) thus ?case by (auto simp: add_nonneg_nonneg)
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-lemma abs_setsum_abs[simp]:
- fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
- shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"
-proof (cases "finite A")
- case True
- thus ?thesis
- proof induct
- case empty thus ?case by simp
- next
- case (insert a A)
- hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp
- also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp
- also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"
- by (simp del: abs_of_nonneg)
- also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp
- finally show ?case .
- qed
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-
-lemma setsum_Plus:
- fixes A :: "'a set" and B :: "'b set"
- assumes fin: "finite A" "finite B"
- shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
-proof -
- have "A <+> B = Inl ` A \<union> Inr ` B" by auto
- moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
- by(auto intro: finite_imageI)
- moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
- moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
- ultimately show ?thesis using fin by(simp add: setsum_Un_disjoint setsum_reindex)
-qed
-
-
-text {* Commuting outer and inner summation *}
-
-lemma swap_inj_on:
- "inj_on (%(i, j). (j, i)) (A \<times> B)"
- by (unfold inj_on_def) fast
-
-lemma swap_product:
- "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
- by (simp add: split_def image_def) blast
-
-lemma setsum_commute:
- "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
-proof (simp add: setsum_cartesian_product)
- have "(\<Sum>(x,y) \<in> A <*> B. f x y) =
- (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"
- (is "?s = _")
- apply (simp add: setsum_reindex [where f = "%(i, j). (j, i)"] swap_inj_on)
- apply (simp add: split_def)
- done
- also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"
- (is "_ = ?t")
- apply (simp add: swap_product)
- done
- finally show "?s = ?t" .
-qed
-
-lemma setsum_product:
- fixes f :: "'a => ('b::semiring_0)"
- shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
- by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
-
-lemma setsum_mult_setsum_if_inj:
-fixes f :: "'a => ('b::semiring_0)"
-shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
- setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
-by(auto simp: setsum_product setsum_cartesian_product
- intro!: setsum_reindex_cong[symmetric])
-
-
-subsection {* Generalized product over a set *}
-
-definition setprod :: "('a => 'b) => 'a set => 'b::comm_monoid_mult"
-where "setprod f A == if finite A then fold_image (op *) f 1 A else 1"
-
-abbreviation
- Setprod ("\<Prod>_" [1000] 999) where
- "\<Prod>A == setprod (%x. x) A"
-
-syntax
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10)
-syntax (xsymbols)
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
-syntax (HTML output)
- "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
-
-translations -- {* Beware of argument permutation! *}
- "PROD i:A. b" == "CONST setprod (%i. b) A"
- "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
-
-text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
- @{text"\<Prod>x|P. e"}. *}
-
-syntax
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
-syntax (xsymbols)
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
-syntax (HTML output)
- "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
-
-translations
- "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
- "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
-
-
-lemma setprod_empty [simp]: "setprod f {} = 1"
-by (auto simp add: setprod_def)
-
-lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==>
- setprod f (insert a A) = f a * setprod f A"
-by (simp add: setprod_def)
-
-lemma setprod_infinite [simp]: "~ finite A ==> setprod f A = 1"
-by (simp add: setprod_def)
-
-lemma setprod_reindex:
- "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
-by(auto simp: setprod_def fold_image_reindex dest!:finite_imageD)
-
-lemma setprod_reindex_id: "inj_on f B ==> setprod f B = setprod id (f ` B)"
-by (auto simp add: setprod_reindex)
-
-lemma setprod_cong:
- "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
-by(fastsimp simp: setprod_def intro: fold_image_cong)
-
-lemma strong_setprod_cong[cong]:
- "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
-by(fastsimp simp: simp_implies_def setprod_def intro: fold_image_cong)
-
-lemma setprod_reindex_cong: "inj_on f A ==>
- B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
-by (frule setprod_reindex, simp)
-
-lemma strong_setprod_reindex_cong: assumes i: "inj_on f A"
- and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
- shows "setprod h B = setprod g A"
-proof-
- have "setprod h B = setprod (h o f) A"
- by (simp add: B setprod_reindex[OF i, of h])
- then show ?thesis apply simp
- apply (rule setprod_cong)
- apply simp
- by (simp add: eq)
-qed
-
-lemma setprod_Un_one:
- assumes fS: "finite S" and fT: "finite T"
- and I0: "\<forall>x \<in> S\<inter>T. f x = 1"
- shows "setprod f (S \<union> T) = setprod f S * setprod f T"
- using fS fT
- apply (simp add: setprod_def)
- apply (rule fold_image_Un_one)
- using I0 by auto
-
-
-lemma setprod_1: "setprod (%i. 1) A = 1"
-apply (case_tac "finite A")
-apply (erule finite_induct, auto simp add: mult_ac)
-done
-
-lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1"
-apply (subgoal_tac "setprod f F = setprod (%x. 1) F")
-apply (erule ssubst, rule setprod_1)
-apply (rule setprod_cong, auto)
-done
-
-lemma setprod_Un_Int: "finite A ==> finite B
- ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
-by(simp add: setprod_def fold_image_Un_Int[symmetric])
-
-lemma setprod_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
-by (subst setprod_Un_Int [symmetric], auto)
-
-lemma setprod_mono_one_left:
- assumes fT: "finite T" and ST: "S \<subseteq> T"
- and z: "\<forall>i \<in> T - S. f i = 1"
- shows "setprod f S = setprod f T"
-proof-
- have eq: "T = S \<union> (T - S)" using ST by blast
- have d: "S \<inter> (T - S) = {}" using ST by blast
- from fT ST have f: "finite S" "finite (T - S)" by (auto intro: finite_subset)
- show ?thesis
- by (simp add: setprod_Un_disjoint[OF f d, unfolded eq[symmetric]] setprod_1'[OF z])
-qed
-
-lemmas setprod_mono_one_right = setprod_mono_one_left [THEN sym]
-
-lemma setprod_delta:
- assumes fS: "finite S"
- shows "setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else 1)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = 1" by simp
- hence ?thesis using a by (simp add: setprod_1 cong add: setprod_cong) }
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have fA1: "setprod ?f ?A = 1" apply (rule setprod_1') by auto
- have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
- using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a by (simp add: fA1 cong add: setprod_cong cong del: if_weak_cong)}
- ultimately show ?thesis by blast
-qed
-
-lemma setprod_delta':
- assumes fS: "finite S" shows
- "setprod (\<lambda>k. if a = k then b k else 1) S =
- (if a\<in> S then b a else 1)"
- using setprod_delta[OF fS, of a b, symmetric]
- by (auto intro: setprod_cong)
-
-
-lemma setprod_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
- setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
-by(simp add: setprod_def fold_image_UN_disjoint cong: setprod_cong)
-
-lemma setprod_Union_disjoint:
- "[| (ALL A:C. finite A);
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) |]
- ==> setprod f (Union C) = setprod (setprod f) C"
-apply (cases "finite C")
- prefer 2 apply (force dest: finite_UnionD simp add: setprod_def)
- apply (frule setprod_UN_disjoint [of C id f])
- apply (unfold Union_def id_def, assumption+)
-done
-
-lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
- (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
- (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
-by(simp add:setprod_def fold_image_Sigma split_def cong:setprod_cong)
-
-text{*Here we can eliminate the finiteness assumptions, by cases.*}
-lemma setprod_cartesian_product:
- "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
-apply (cases "finite A")
- apply (cases "finite B")
- apply (simp add: setprod_Sigma)
- apply (cases "A={}", simp)
- apply (simp add: setprod_1)
-apply (auto simp add: setprod_def
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma setprod_timesf:
- "setprod (%x. f x * g x) A = (setprod f A * setprod g A)"
-by(simp add:setprod_def fold_image_distrib)
-
-
-subsubsection {* Properties in more restricted classes of structures *}
-
-lemma setprod_eq_1_iff [simp]:
- "finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))"
-by (induct set: finite) auto
-
-lemma setprod_zero:
- "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"
-apply (induct set: finite, force, clarsimp)
-apply (erule disjE, auto)
-done
-
-lemma setprod_nonneg [rule_format]:
- "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
-by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)
-
-lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)
- --> 0 < setprod f A"
-by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)
-
-lemma setprod_zero_iff[simp]: "finite A ==>
- (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
- (EX x: A. f x = 0)"
-by (erule finite_induct, auto simp:no_zero_divisors)
-
-lemma setprod_pos_nat:
- "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
-using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
-
-lemma setprod_pos_nat_iff[simp]:
- "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"
-using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
-
-lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
- (setprod f (A Un B) :: 'a ::{field})
- = setprod f A * setprod f B / setprod f (A Int B)"
-by (subst setprod_Un_Int [symmetric], auto)
-
-lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
- (setprod f (A - {a}) :: 'a :: {field}) =
- (if a:A then setprod f A / f a else setprod f A)"
-by (erule finite_induct) (auto simp add: insert_Diff_if)
-
-lemma setprod_inversef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
- shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
-by (erule finite_induct) auto
-
-lemma setprod_dividef:
- fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
- shows "finite A
- ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
-apply (subgoal_tac
- "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
-apply (erule ssubst)
-apply (subst divide_inverse)
-apply (subst setprod_timesf)
-apply (subst setprod_inversef, assumption+, rule refl)
-apply (rule setprod_cong, rule refl)
-apply (subst divide_inverse, auto)
-done
-
-lemma setprod_dvd_setprod [rule_format]:
- "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"
- apply (cases "finite A")
- apply (induct set: finite)
- apply (auto simp add: dvd_def)
- apply (rule_tac x = "k * ka" in exI)
- apply (simp add: algebra_simps)
-done
-
-lemma setprod_dvd_setprod_subset:
- "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
- apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
- apply (unfold dvd_def, blast)
- apply (subst setprod_Un_disjoint [symmetric])
- apply (auto elim: finite_subset intro: setprod_cong)
-done
-
-lemma setprod_dvd_setprod_subset2:
- "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>
- setprod f A dvd setprod g B"
- apply (rule dvd_trans)
- apply (rule setprod_dvd_setprod, erule (1) bspec)
- apply (erule (1) setprod_dvd_setprod_subset)
-done
-
-lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>
- (f i ::'a::comm_semiring_1) dvd setprod f A"
-by (induct set: finite) (auto intro: dvd_mult)
-
-lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>
- (d::'a::comm_semiring_1) dvd (SUM x : A. f x)"
- apply (cases "finite A")
- apply (induct set: finite)
- apply auto
-done
-
-lemma setprod_mono:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"
- assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"
- shows "setprod f A \<le> setprod g A"
-proof (cases "finite A")
- case True
- hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]
- proof (induct A rule: finite_subset_induct)
- case (insert a F)
- thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
- unfolding setprod_insert[OF insert(1,3)]
- using assms[rule_format,OF insert(2)] insert
- by (auto intro: mult_mono mult_nonneg_nonneg)
- qed auto
- thus ?thesis by simp
-qed auto
-
-lemma abs_setprod:
- fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
- shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"
-proof (cases "finite A")
- case True thus ?thesis
- by induct (auto simp add: field_simps abs_mult)
-qed auto
-
-
-subsection {* Finite cardinality *}
-
-text {* This definition, although traditional, is ugly to work with:
-@{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}.
-But now that we have @{text setsum} things are easy:
-*}
-
-definition card :: "'a set \<Rightarrow> nat" where
- "card A = setsum (\<lambda>x. 1) A"
-
-lemmas card_eq_setsum = card_def
-
-lemma card_empty [simp]: "card {} = 0"
- by (simp add: card_def)
-
-lemma card_insert_disjoint [simp]:
- "finite A ==> x \<notin> A ==> card (insert x A) = Suc(card A)"
- by (simp add: card_def)
-
-lemma card_insert_if:
- "finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))"
- by (simp add: insert_absorb)
-
-lemma card_infinite [simp]: "~ finite A ==> card A = 0"
- by (simp add: card_def)
-
-lemma card_ge_0_finite:
- "card A > 0 \<Longrightarrow> finite A"
- by (rule ccontr) simp
-
-lemma card_0_eq [simp,noatp]: "finite A ==> (card A = 0) = (A = {})"
- apply auto
- apply (drule_tac a = x in mk_disjoint_insert, clarify, auto)
- done
-
-lemma finite_UNIV_card_ge_0:
- "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
- by (rule ccontr) simp
-
-lemma card_eq_0_iff: "(card A = 0) = (A = {} | ~ finite A)"
- by auto
-
-lemma card_gt_0_iff: "(0 < card A) = (A \<noteq> {} & finite A)"
- by (simp add: neq0_conv [symmetric] card_eq_0_iff)
-
-lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
-apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
-apply(simp del:insert_Diff_single)
-done
-
-lemma card_Diff_singleton:
- "finite A ==> x: A ==> card (A - {x}) = card A - 1"
-by (simp add: card_Suc_Diff1 [symmetric])
-
-lemma card_Diff_singleton_if:
- "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
-by (simp add: card_Diff_singleton)
-
-lemma card_Diff_insert[simp]:
-assumes "finite A" and "a:A" and "a ~: B"
-shows "card(A - insert a B) = card(A - B) - 1"
-proof -
- have "A - insert a B = (A - B) - {a}" using assms by blast
- then show ?thesis using assms by(simp add:card_Diff_singleton)
-qed
-
-lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
-by (simp add: card_insert_if card_Suc_Diff1 del:card_Diff_insert)
-
-lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
-by (simp add: card_insert_if)
-
-lemma card_mono: "\<lbrakk> finite B; A \<subseteq> B \<rbrakk> \<Longrightarrow> card A \<le> card B"
-by (simp add: card_def setsum_mono2)
-
-lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
-apply (induct set: finite, simp, clarify)
-apply (subgoal_tac "finite A & A - {x} <= F")
- prefer 2 apply (blast intro: finite_subset, atomize)
-apply (drule_tac x = "A - {x}" in spec)
-apply (simp add: card_Diff_singleton_if split add: split_if_asm)
-apply (case_tac "card A", auto)
-done
-
-lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
-apply (simp add: psubset_eq linorder_not_le [symmetric])
-apply (blast dest: card_seteq)
-done
-
-lemma card_Un_Int: "finite A ==> finite B
- ==> card A + card B = card (A Un B) + card (A Int B)"
-by(simp add:card_def setsum_Un_Int)
-
-lemma card_Un_disjoint: "finite A ==> finite B
- ==> A Int B = {} ==> card (A Un B) = card A + card B"
-by (simp add: card_Un_Int)
-
-lemma card_Diff_subset:
- "finite B ==> B <= A ==> card (A - B) = card A - card B"
-by(simp add:card_def setsum_diff_nat)
-
-lemma card_Diff_subset_Int:
- assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
-proof -
- have "A - B = A - A \<inter> B" by auto
- thus ?thesis
- by (simp add: card_Diff_subset AB)
-qed
-
-lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
-apply (rule Suc_less_SucD)
-apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
-done
-
-lemma card_Diff2_less:
- "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
-apply (case_tac "x = y")
- apply (simp add: card_Diff1_less del:card_Diff_insert)
-apply (rule less_trans)
- prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
-done
-
-lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
-apply (case_tac "x : A")
- apply (simp_all add: card_Diff1_less less_imp_le)
-done
-
-lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
-by (erule psubsetI, blast)
-
-lemma insert_partition:
- "\<lbrakk> x \<notin> F; \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<rbrakk>
- \<Longrightarrow> x \<inter> \<Union> F = {}"
-by auto
-
-lemma finite_psubset_induct[consumes 1, case_names psubset]:
- assumes "finite A" and "!!A. finite A \<Longrightarrow> (!!B. finite B \<Longrightarrow> B \<subset> A \<Longrightarrow> P(B)) \<Longrightarrow> P(A)" shows "P A"
-using assms(1)
-proof (induct A rule: measure_induct_rule[where f=card])
- case (less A)
- show ?case
- proof(rule assms(2)[OF less(2)])
- fix B assume "finite B" "B \<subset> A"
- show "P B" by(rule less(1)[OF psubset_card_mono[OF less(2) `B \<subset> A`] `finite B`])
- qed
-qed
-
-text{* main cardinality theorem *}
-lemma card_partition [rule_format]:
- "finite C ==>
- finite (\<Union> C) -->
- (\<forall>c\<in>C. card c = k) -->
- (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
- k * card(C) = card (\<Union> C)"
-apply (erule finite_induct, simp)
-apply (simp add: card_Un_disjoint insert_partition
- finite_subset [of _ "\<Union> (insert x F)"])
-done
-
-lemma card_eq_UNIV_imp_eq_UNIV:
- assumes fin: "finite (UNIV :: 'a set)"
- and card: "card A = card (UNIV :: 'a set)"
- shows "A = (UNIV :: 'a set)"
-proof
- show "A \<subseteq> UNIV" by simp
- show "UNIV \<subseteq> A"
- proof
- fix x
- show "x \<in> A"
- proof (rule ccontr)
- assume "x \<notin> A"
- then have "A \<subset> UNIV" by auto
- with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
- with card show False by simp
- qed
- qed
-qed
-
-text{*The form of a finite set of given cardinality*}
-
-lemma card_eq_SucD:
-assumes "card A = Suc k"
-shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
-proof -
- have fin: "finite A" using assms by (auto intro: ccontr)
- moreover have "card A \<noteq> 0" using assms by auto
- ultimately obtain b where b: "b \<in> A" by auto
- show ?thesis
- proof (intro exI conjI)
- show "A = insert b (A-{b})" using b by blast
- show "b \<notin> A - {b}" by blast
- show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
- using assms b fin by(fastsimp dest:mk_disjoint_insert)+
- qed
-qed
-
-lemma card_Suc_eq:
- "(card A = Suc k) =
- (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
-apply(rule iffI)
- apply(erule card_eq_SucD)
-apply(auto)
-apply(subst card_insert)
- apply(auto intro:ccontr)
-done
-
-lemma finite_fun_UNIVD2:
- assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
- shows "finite (UNIV :: 'b set)"
-proof -
- from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
- by(rule finite_imageI)
- moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
- by(rule UNIV_eq_I) auto
- ultimately show "finite (UNIV :: 'b set)" by simp
-qed
-
-lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat(card A) * y"
-apply (cases "finite A")
-apply (erule finite_induct)
-apply (auto simp add: algebra_simps)
-done
-
-lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
-apply (erule finite_induct)
-apply auto
-done
-
-lemma setprod_gen_delta:
- assumes fS: "finite S"
- shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::{comm_monoid_mult}) * c^ (card S - 1) else c^ card S)"
-proof-
- let ?f = "(\<lambda>k. if k=a then b k else c)"
- {assume a: "a \<notin> S"
- hence "\<forall> k\<in> S. ?f k = c" by simp
- hence ?thesis using a setprod_constant[OF fS, of c] by (simp add: setprod_1 cong add: setprod_cong) }
- moreover
- {assume a: "a \<in> S"
- let ?A = "S - {a}"
- let ?B = "{a}"
- have eq: "S = ?A \<union> ?B" using a by blast
- have dj: "?A \<inter> ?B = {}" by simp
- from fS have fAB: "finite ?A" "finite ?B" by auto
- have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
- apply (rule setprod_cong) by auto
- have cA: "card ?A = card S - 1" using fS a by auto
- have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto
- have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
- using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
- by simp
- then have ?thesis using a cA
- by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
- ultimately show ?thesis by blast
-qed
-
-
-lemma setsum_bounded:
- assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
- shows "setsum f A \<le> of_nat(card A) * K"
-proof (cases "finite A")
- case True
- thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp
-next
- case False thus ?thesis by (simp add: setsum_def)
-qed
-
-
-lemma card_UNIV_unit: "card (UNIV :: unit set) = 1"
- unfolding UNIV_unit by simp
-
-
-subsubsection {* Cardinality of unions *}
-
-lemma card_UN_disjoint:
- "finite I ==> (ALL i:I. finite (A i)) ==>
- (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {})
- ==> card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
-apply (simp add: card_def del: setsum_constant)
-apply (subgoal_tac
- "setsum (%i. card (A i)) I = setsum (%i. (setsum (%x. 1) (A i))) I")
-apply (simp add: setsum_UN_disjoint del: setsum_constant)
-apply (simp cong: setsum_cong)
-done
-
-lemma card_Union_disjoint:
- "finite C ==> (ALL A:C. finite A) ==>
- (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
- ==> card (Union C) = setsum card C"
-apply (frule card_UN_disjoint [of C id])
-apply (unfold Union_def id_def, assumption+)
-done
-
-
-subsubsection {* Cardinality of image *}
-
-text{*The image of a finite set can be expressed using @{term fold_image}.*}
-lemma image_eq_fold_image:
- "finite A ==> f ` A = fold_image (op Un) (%x. {f x}) {} A"
-proof (induct rule: finite_induct)
- case empty then show ?case by simp
-next
- interpret ab_semigroup_mult "op Un"
- proof qed auto
- case insert
- then show ?case by simp
-qed
-
-lemma card_image_le: "finite A ==> card (f ` A) <= card A"
-apply (induct set: finite)
- apply simp
-apply (simp add: le_SucI card_insert_if)
-done
-
-lemma card_image: "inj_on f A ==> card (f ` A) = card A"
-by(simp add:card_def setsum_reindex o_def del:setsum_constant)
-
-lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
-by(auto simp: card_image bij_betw_def)
-
-lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
-by (simp add: card_seteq card_image)
-
-lemma eq_card_imp_inj_on:
- "[| finite A; card(f ` A) = card A |] ==> inj_on f A"
-apply (induct rule:finite_induct)
-apply simp
-apply(frule card_image_le[where f = f])
-apply(simp add:card_insert_if split:if_splits)
-done
-
-lemma inj_on_iff_eq_card:
- "finite A ==> inj_on f A = (card(f ` A) = card A)"
-by(blast intro: card_image eq_card_imp_inj_on)
-
-
-lemma card_inj_on_le:
- "[|inj_on f A; f ` A \<subseteq> B; finite B |] ==> card A \<le> card B"
-apply (subgoal_tac "finite A")
- apply (force intro: card_mono simp add: card_image [symmetric])
-apply (blast intro: finite_imageD dest: finite_subset)
-done
-
-lemma card_bij_eq:
- "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
- finite A; finite B |] ==> card A = card B"
-by (auto intro: le_antisym card_inj_on_le)
-
-
-subsubsection {* Cardinality of products *}
-
-(*
-lemma SigmaI_insert: "y \<notin> A ==>
- (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
- by auto
-*)
-
-lemma card_SigmaI [simp]:
- "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
- \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
-by(simp add:card_def setsum_Sigma del:setsum_constant)
-
-lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"
-apply (cases "finite A")
-apply (cases "finite B")
-apply (auto simp add: card_eq_0_iff
- dest: finite_cartesian_productD1 finite_cartesian_productD2)
-done
-
-lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)"
-by (simp add: card_cartesian_product)
-
-
-subsubsection {* Cardinality of sums *}
-
-lemma card_Plus:
- assumes "finite A" and "finite B"
- shows "card (A <+> B) = card A + card B"
-proof -
- have "Inl`A \<inter> Inr`B = {}" by fast
- with assms show ?thesis
- unfolding Plus_def
- by (simp add: card_Un_disjoint card_image)
-qed
-
-lemma card_Plus_conv_if:
- "card (A <+> B) = (if finite A \<and> finite B then card(A) + card(B) else 0)"
-by(auto simp: card_def setsum_Plus simp del: setsum_constant)
-
-
-subsubsection {* Cardinality of the Powerset *}
-
-lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *)
-apply (induct set: finite)
- apply (simp_all add: Pow_insert)
-apply (subst card_Un_disjoint, blast)
- apply (blast intro: finite_imageI, blast)
-apply (subgoal_tac "inj_on (insert x) (Pow F)")
- apply (simp add: card_image Pow_insert)
-apply (unfold inj_on_def)
-apply (blast elim!: equalityE)
-done
-
-text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
-
-lemma dvd_partition:
- "finite (Union C) ==>
- ALL c : C. k dvd card c ==>
- (ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
- k dvd card (Union C)"
-apply(frule finite_UnionD)
-apply(rotate_tac -1)
-apply (induct set: finite, simp_all, clarify)
-apply (subst card_Un_disjoint)
- apply (auto simp add: disjoint_eq_subset_Compl)
-done
-
-
-subsubsection {* Relating injectivity and surjectivity *}
-
-lemma finite_surj_inj: "finite(A) \<Longrightarrow> A <= f`A \<Longrightarrow> inj_on f A"
-apply(rule eq_card_imp_inj_on, assumption)
-apply(frule finite_imageI)
-apply(drule (1) card_seteq)
- apply(erule card_image_le)
-apply simp
-done
-
-lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
-shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
-by (blast intro: finite_surj_inj subset_UNIV dest:surj_range)
-
-lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
-shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
-by(fastsimp simp:surj_def dest!: endo_inj_surj)
-
-corollary infinite_UNIV_nat[iff]: "~finite(UNIV::nat set)"
-proof
- assume "finite(UNIV::nat set)"
- with finite_UNIV_inj_surj[of Suc]
- show False by simp (blast dest: Suc_neq_Zero surjD)
-qed
-
-(* Often leads to bogus ATP proofs because of reduced type information, hence noatp *)
-lemma infinite_UNIV_char_0[noatp]:
- "\<not> finite (UNIV::'a::semiring_char_0 set)"
-proof
- assume "finite (UNIV::'a set)"
- with subset_UNIV have "finite (range of_nat::'a set)"
- by (rule finite_subset)
- moreover have "inj (of_nat::nat \<Rightarrow> 'a)"
- by (simp add: inj_on_def)
- ultimately have "finite (UNIV::nat set)"
- by (rule finite_imageD)
- then show "False"
- by simp
-qed
subsection{* A fold functional for non-empty sets *}
@@ -2811,561 +1354,6 @@
qed
-subsubsection {* Fold1 in lattices with @{const inf} and @{const sup} *}
-
-text{*
- As an application of @{text fold1} we define infimum
- and supremum in (not necessarily complete!) lattices
- over (non-empty) sets by means of @{text fold1}.
-*}
-
-context semilattice_inf
-begin
-
-lemma below_fold1_iff:
- assumes "finite A" "A \<noteq> {}"
- shows "x \<le> fold1 inf A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- show ?thesis using assms by (induct rule: finite_ne_induct) simp_all
-qed
-
-lemma fold1_belowI:
- assumes "finite A"
- and "a \<in> A"
- shows "fold1 inf A \<le> a"
-proof -
- from assms have "A \<noteq> {}" by auto
- from `finite A` `A \<noteq> {}` `a \<in> A` show ?thesis
- proof (induct rule: finite_ne_induct)
- case singleton thus ?case by simp
- next
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- case (insert x F)
- from insert(5) have "a = x \<or> a \<in> F" by simp
- thus ?case
- proof
- assume "a = x" thus ?thesis using insert
- by (simp add: mult_ac)
- next
- assume "a \<in> F"
- hence bel: "fold1 inf F \<le> a" by (rule insert)
- have "inf (fold1 inf (insert x F)) a = inf x (inf (fold1 inf F) a)"
- using insert by (simp add: mult_ac)
- also have "inf (fold1 inf F) a = fold1 inf F"
- using bel by (auto intro: antisym)
- also have "inf x \<dots> = fold1 inf (insert x F)"
- using insert by (simp add: mult_ac)
- finally have aux: "inf (fold1 inf (insert x F)) a = fold1 inf (insert x F)" .
- moreover have "inf (fold1 inf (insert x F)) a \<le> a" by simp
- ultimately show ?thesis by simp
- qed
- qed
-qed
-
-end
-
-context lattice
-begin
-
-definition
- Inf_fin :: "'a set \<Rightarrow> 'a" ("\<Sqinter>\<^bsub>fin\<^esub>_" [900] 900)
-where
- "Inf_fin = fold1 inf"
-
-definition
- Sup_fin :: "'a set \<Rightarrow> 'a" ("\<Squnion>\<^bsub>fin\<^esub>_" [900] 900)
-where
- "Sup_fin = fold1 sup"
-
-lemma Inf_le_Sup [simp]: "\<lbrakk> finite A; A \<noteq> {} \<rbrakk> \<Longrightarrow> \<Sqinter>\<^bsub>fin\<^esub>A \<le> \<Squnion>\<^bsub>fin\<^esub>A"
-apply(unfold Sup_fin_def Inf_fin_def)
-apply(subgoal_tac "EX a. a:A")
-prefer 2 apply blast
-apply(erule exE)
-apply(rule order_trans)
-apply(erule (1) fold1_belowI)
-apply(erule (1) semilattice_inf.fold1_belowI [OF dual_semilattice])
-done
-
-lemma sup_Inf_absorb [simp]:
- "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> sup a (\<Sqinter>\<^bsub>fin\<^esub>A) = a"
-apply(subst sup_commute)
-apply(simp add: Inf_fin_def sup_absorb2 fold1_belowI)
-done
-
-lemma inf_Sup_absorb [simp]:
- "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> inf a (\<Squnion>\<^bsub>fin\<^esub>A) = a"
-by (simp add: Sup_fin_def inf_absorb1
- semilattice_inf.fold1_belowI [OF dual_semilattice])
-
-end
-
-context distrib_lattice
-begin
-
-lemma sup_Inf1_distrib:
- assumes "finite A"
- and "A \<noteq> {}"
- shows "sup x (\<Sqinter>\<^bsub>fin\<^esub>A) = \<Sqinter>\<^bsub>fin\<^esub>{sup x a|a. a \<in> A}"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- from assms show ?thesis
- by (simp add: Inf_fin_def image_def
- hom_fold1_commute [where h="sup x", OF sup_inf_distrib1])
- (rule arg_cong [where f="fold1 inf"], blast)
-qed
-
-lemma sup_Inf2_distrib:
- assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
- shows "sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B) = \<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B}"
-using A proof (induct rule: finite_ne_induct)
- case singleton thus ?case
- by (simp add: sup_Inf1_distrib [OF B] fold1_singleton_def [OF Inf_fin_def])
-next
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- case (insert x A)
- have finB: "finite {sup x b |b. b \<in> B}"
- by(rule finite_surj[where f = "sup x", OF B(1)], auto)
- have finAB: "finite {sup a b |a b. a \<in> A \<and> b \<in> B}"
- proof -
- have "{sup a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {sup a b})"
- by blast
- thus ?thesis by(simp add: insert(1) B(1))
- qed
- have ne: "{sup a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
- have "sup (\<Sqinter>\<^bsub>fin\<^esub>(insert x A)) (\<Sqinter>\<^bsub>fin\<^esub>B) = sup (inf x (\<Sqinter>\<^bsub>fin\<^esub>A)) (\<Sqinter>\<^bsub>fin\<^esub>B)"
- using insert by (simp add: fold1_insert_idem_def [OF Inf_fin_def])
- also have "\<dots> = inf (sup x (\<Sqinter>\<^bsub>fin\<^esub>B)) (sup (\<Sqinter>\<^bsub>fin\<^esub>A) (\<Sqinter>\<^bsub>fin\<^esub>B))" by(rule sup_inf_distrib2)
- also have "\<dots> = inf (\<Sqinter>\<^bsub>fin\<^esub>{sup x b|b. b \<in> B}) (\<Sqinter>\<^bsub>fin\<^esub>{sup a b|a b. a \<in> A \<and> b \<in> B})"
- using insert by(simp add:sup_Inf1_distrib[OF B])
- also have "\<dots> = \<Sqinter>\<^bsub>fin\<^esub>({sup x b |b. b \<in> B} \<union> {sup a b |a b. a \<in> A \<and> b \<in> B})"
- (is "_ = \<Sqinter>\<^bsub>fin\<^esub>?M")
- using B insert
- by (simp add: Inf_fin_def fold1_Un2 [OF finB _ finAB ne])
- also have "?M = {sup a b |a b. a \<in> insert x A \<and> b \<in> B}"
- by blast
- finally show ?case .
-qed
-
-lemma inf_Sup1_distrib:
- assumes "finite A" and "A \<noteq> {}"
- shows "inf x (\<Squnion>\<^bsub>fin\<^esub>A) = \<Squnion>\<^bsub>fin\<^esub>{inf x a|a. a \<in> A}"
-proof -
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- from assms show ?thesis
- by (simp add: Sup_fin_def image_def hom_fold1_commute [where h="inf x", OF inf_sup_distrib1])
- (rule arg_cong [where f="fold1 sup"], blast)
-qed
-
-lemma inf_Sup2_distrib:
- assumes A: "finite A" "A \<noteq> {}" and B: "finite B" "B \<noteq> {}"
- shows "inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B) = \<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B}"
-using A proof (induct rule: finite_ne_induct)
- case singleton thus ?case
- by(simp add: inf_Sup1_distrib [OF B] fold1_singleton_def [OF Sup_fin_def])
-next
- case (insert x A)
- have finB: "finite {inf x b |b. b \<in> B}"
- by(rule finite_surj[where f = "%b. inf x b", OF B(1)], auto)
- have finAB: "finite {inf a b |a b. a \<in> A \<and> b \<in> B}"
- proof -
- have "{inf a b |a b. a \<in> A \<and> b \<in> B} = (UN a:A. UN b:B. {inf a b})"
- by blast
- thus ?thesis by(simp add: insert(1) B(1))
- qed
- have ne: "{inf a b |a b. a \<in> A \<and> b \<in> B} \<noteq> {}" using insert B by blast
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- have "inf (\<Squnion>\<^bsub>fin\<^esub>(insert x A)) (\<Squnion>\<^bsub>fin\<^esub>B) = inf (sup x (\<Squnion>\<^bsub>fin\<^esub>A)) (\<Squnion>\<^bsub>fin\<^esub>B)"
- using insert by (simp add: fold1_insert_idem_def [OF Sup_fin_def])
- also have "\<dots> = sup (inf x (\<Squnion>\<^bsub>fin\<^esub>B)) (inf (\<Squnion>\<^bsub>fin\<^esub>A) (\<Squnion>\<^bsub>fin\<^esub>B))" by(rule inf_sup_distrib2)
- also have "\<dots> = sup (\<Squnion>\<^bsub>fin\<^esub>{inf x b|b. b \<in> B}) (\<Squnion>\<^bsub>fin\<^esub>{inf a b|a b. a \<in> A \<and> b \<in> B})"
- using insert by(simp add:inf_Sup1_distrib[OF B])
- also have "\<dots> = \<Squnion>\<^bsub>fin\<^esub>({inf x b |b. b \<in> B} \<union> {inf a b |a b. a \<in> A \<and> b \<in> B})"
- (is "_ = \<Squnion>\<^bsub>fin\<^esub>?M")
- using B insert
- by (simp add: Sup_fin_def fold1_Un2 [OF finB _ finAB ne])
- also have "?M = {inf a b |a b. a \<in> insert x A \<and> b \<in> B}"
- by blast
- finally show ?case .
-qed
-
-end
-
-
-subsubsection {* Fold1 in linear orders with @{const min} and @{const max} *}
-
-text{*
- As an application of @{text fold1} we define minimum
- and maximum in (not necessarily complete!) linear orders
- over (non-empty) sets by means of @{text fold1}.
-*}
-
-context linorder
-begin
-
-lemma ab_semigroup_idem_mult_min:
- "ab_semigroup_idem_mult min"
- proof qed (auto simp add: min_def)
-
-lemma ab_semigroup_idem_mult_max:
- "ab_semigroup_idem_mult max"
- proof qed (auto simp add: max_def)
-
-lemma max_lattice:
- "semilattice_inf (op \<ge>) (op >) max"
- by (fact min_max.dual_semilattice)
-
-lemma dual_max:
- "ord.max (op \<ge>) = min"
- by (auto simp add: ord.max_def_raw min_def expand_fun_eq)
-
-lemma dual_min:
- "ord.min (op \<ge>) = max"
- by (auto simp add: ord.min_def_raw max_def expand_fun_eq)
-
-lemma strict_below_fold1_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < fold1 min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert)
-qed
-
-lemma fold1_below_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "fold1 min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert min_le_iff_disj)
-qed
-
-lemma fold1_strict_below_iff:
- assumes "finite A" and "A \<noteq> {}"
- shows "fold1 min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (induct rule: finite_ne_induct)
- (simp_all add: fold1_insert min_less_iff_disj)
-qed
-
-lemma fold1_antimono:
- assumes "A \<noteq> {}" and "A \<subseteq> B" and "finite B"
- shows "fold1 min B \<le> fold1 min A"
-proof cases
- assume "A = B" thus ?thesis by simp
-next
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- assume "A \<noteq> B"
- have B: "B = A \<union> (B-A)" using `A \<subseteq> B` by blast
- have "fold1 min B = fold1 min (A \<union> (B-A))" by(subst B)(rule refl)
- also have "\<dots> = min (fold1 min A) (fold1 min (B-A))"
- proof -
- have "finite A" by(rule finite_subset[OF `A \<subseteq> B` `finite B`])
- moreover have "finite(B-A)" by(rule finite_Diff[OF `finite B`]) (* by(blast intro:finite_Diff prems) fails *)
- moreover have "(B-A) \<noteq> {}" using prems by blast
- moreover have "A Int (B-A) = {}" using prems by blast
- ultimately show ?thesis using `A \<noteq> {}` by (rule_tac fold1_Un)
- qed
- also have "\<dots> \<le> fold1 min A" by (simp add: min_le_iff_disj)
- finally show ?thesis .
-qed
-
-definition
- Min :: "'a set \<Rightarrow> 'a"
-where
- "Min = fold1 min"
-
-definition
- Max :: "'a set \<Rightarrow> 'a"
-where
- "Max = fold1 max"
-
-lemmas Min_singleton [simp] = fold1_singleton_def [OF Min_def]
-lemmas Max_singleton [simp] = fold1_singleton_def [OF Max_def]
-
-lemma Min_insert [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min (insert x A) = min x (Min A)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis by (rule fold1_insert_idem_def [OF Min_def])
-qed
-
-lemma Max_insert [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max (insert x A) = max x (Max A)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis by (rule fold1_insert_idem_def [OF Max_def])
-qed
-
-lemma Min_in [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A \<in> A"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms fold1_in show ?thesis by (fastsimp simp: Min_def min_def)
-qed
-
-lemma Max_in [simp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A \<in> A"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms fold1_in [of A] show ?thesis by (fastsimp simp: Max_def max_def)
-qed
-
-lemma Min_Un:
- assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
- shows "Min (A \<union> B) = min (Min A) (Min B)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (simp add: Min_def fold1_Un2)
-qed
-
-lemma Max_Un:
- assumes "finite A" and "A \<noteq> {}" and "finite B" and "B \<noteq> {}"
- shows "Max (A \<union> B) = max (Max A) (Max B)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis
- by (simp add: Max_def fold1_Un2)
-qed
-
-lemma hom_Min_commute:
- assumes "\<And>x y. h (min x y) = min (h x) (h y)"
- and "finite N" and "N \<noteq> {}"
- shows "h (Min N) = Min (h ` N)"
-proof -
- interpret ab_semigroup_idem_mult min
- by (rule ab_semigroup_idem_mult_min)
- from assms show ?thesis
- by (simp add: Min_def hom_fold1_commute)
-qed
-
-lemma hom_Max_commute:
- assumes "\<And>x y. h (max x y) = max (h x) (h y)"
- and "finite N" and "N \<noteq> {}"
- shows "h (Max N) = Max (h ` N)"
-proof -
- interpret ab_semigroup_idem_mult max
- by (rule ab_semigroup_idem_mult_max)
- from assms show ?thesis
- by (simp add: Max_def hom_fold1_commute [of h])
-qed
-
-lemma Min_le [simp]:
- assumes "finite A" and "x \<in> A"
- shows "Min A \<le> x"
- using assms by (simp add: Min_def min_max.fold1_belowI)
-
-lemma Max_ge [simp]:
- assumes "finite A" and "x \<in> A"
- shows "x \<le> Max A"
-proof -
- interpret semilattice_inf "op \<ge>" "op >" max
- by (rule max_lattice)
- from assms show ?thesis by (simp add: Max_def fold1_belowI)
-qed
-
-lemma Min_ge_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
- using assms by (simp add: Min_def min_max.below_fold1_iff)
-
-lemma Max_le_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A \<le> x \<longleftrightarrow> (\<forall>a\<in>A. a \<le> x)"
-proof -
- interpret semilattice_inf "op \<ge>" "op >" max
- by (rule max_lattice)
- from assms show ?thesis by (simp add: Max_def below_fold1_iff)
-qed
-
-lemma Min_gr_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
- using assms by (simp add: Min_def strict_below_fold1_iff)
-
-lemma Max_less_iff [simp, noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.strict_below_fold1_iff [folded dual.dual_max])
-qed
-
-lemma Min_le_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
- using assms by (simp add: Min_def fold1_below_iff)
-
-lemma Max_ge_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_below_iff [folded dual.dual_max])
-qed
-
-lemma Min_less_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
- using assms by (simp add: Min_def fold1_strict_below_iff)
-
-lemma Max_gr_iff [noatp]:
- assumes "finite A" and "A \<noteq> {}"
- shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_strict_below_iff [folded dual.dual_max])
-qed
-
-lemma Min_eqI:
- assumes "finite A"
- assumes "\<And>y. y \<in> A \<Longrightarrow> y \<ge> x"
- and "x \<in> A"
- shows "Min A = x"
-proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
- with assms show "Min A \<ge> x" by simp
-next
- from assms show "x \<ge> Min A" by simp
-qed
-
-lemma Max_eqI:
- assumes "finite A"
- assumes "\<And>y. y \<in> A \<Longrightarrow> y \<le> x"
- and "x \<in> A"
- shows "Max A = x"
-proof (rule antisym)
- from `x \<in> A` have "A \<noteq> {}" by auto
- with assms show "Max A \<le> x" by simp
-next
- from assms show "x \<le> Max A" by simp
-qed
-
-lemma Min_antimono:
- assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
- shows "Min N \<le> Min M"
- using assms by (simp add: Min_def fold1_antimono)
-
-lemma Max_mono:
- assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
- shows "Max M \<le> Max N"
-proof -
- interpret dual: linorder "op \<ge>" "op >"
- by (rule dual_linorder)
- from assms show ?thesis
- by (simp add: Max_def dual.fold1_antimono [folded dual.dual_max])
-qed
-
-lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:
- "finite A \<Longrightarrow> P {} \<Longrightarrow>
- (!!b A. finite A \<Longrightarrow> ALL a:A. a < b \<Longrightarrow> P A \<Longrightarrow> P(insert b A))
- \<Longrightarrow> P A"
-proof (induct rule: finite_psubset_induct)
- fix A :: "'a set"
- assume IH: "!! B. finite B \<Longrightarrow> B < A \<Longrightarrow> P {} \<Longrightarrow>
- (!!b A. finite A \<Longrightarrow> (\<forall>a\<in>A. a<b) \<Longrightarrow> P A \<Longrightarrow> P (insert b A))
- \<Longrightarrow> P B"
- and "finite A" and "P {}"
- and step: "!!b A. \<lbrakk>finite A; \<forall>a\<in>A. a < b; P A\<rbrakk> \<Longrightarrow> P (insert b A)"
- show "P A"
- proof (cases "A = {}")
- assume "A = {}" thus "P A" using `P {}` by simp
- next
- let ?B = "A - {Max A}" let ?A = "insert (Max A) ?B"
- assume "A \<noteq> {}"
- with `finite A` have "Max A : A" by auto
- hence A: "?A = A" using insert_Diff_single insert_absorb by auto
- moreover have "finite ?B" using `finite A` by simp
- ultimately have "P ?B" using `P {}` step IH[of ?B] by blast
- moreover have "\<forall>a\<in>?B. a < Max A" using Max_ge [OF `finite A`] by fastsimp
- ultimately show "P A" using A insert_Diff_single step[OF `finite ?B`] by fastsimp
- qed
-qed
-
-lemma finite_linorder_min_induct[consumes 1, case_names empty insert]:
- "\<lbrakk>finite A; P {}; \<And>b A. \<lbrakk>finite A; \<forall>a\<in>A. b < a; P A\<rbrakk> \<Longrightarrow> P (insert b A)\<rbrakk> \<Longrightarrow> P A"
-by(rule linorder.finite_linorder_max_induct[OF dual_linorder])
-
-end
-
-context linordered_ab_semigroup_add
-begin
-
-lemma add_Min_commute:
- fixes k
- assumes "finite N" and "N \<noteq> {}"
- shows "k + Min N = Min {k + m | m. m \<in> N}"
-proof -
- have "\<And>x y. k + min x y = min (k + x) (k + y)"
- by (simp add: min_def not_le)
- (blast intro: antisym less_imp_le add_left_mono)
- with assms show ?thesis
- using hom_Min_commute [of "plus k" N]
- by simp (blast intro: arg_cong [where f = Min])
-qed
-
-lemma add_Max_commute:
- fixes k
- assumes "finite N" and "N \<noteq> {}"
- shows "k + Max N = Max {k + m | m. m \<in> N}"
-proof -
- have "\<And>x y. k + max x y = max (k + x) (k + y)"
- by (simp add: max_def not_le)
- (blast intro: antisym less_imp_le add_left_mono)
- with assms show ?thesis
- using hom_Max_commute [of "plus k" N]
- by simp (blast intro: arg_cong [where f = Max])
-qed
-
-end
-
-context linordered_ab_group_add
-begin
-
-lemma minus_Max_eq_Min [simp]:
- "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Max S) = Min (uminus ` S)"
- by (induct S rule: finite_ne_induct) (simp_all add: minus_max_eq_min)
-
-lemma minus_Min_eq_Max [simp]:
- "finite S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> - (Min S) = Max (uminus ` S)"
- by (induct S rule: finite_ne_induct) (simp_all add: minus_min_eq_max)
-
-end
-
-
subsection {* Expressing set operations via @{const fold} *}
lemma (in fun_left_comm) fun_left_comm_apply:
@@ -3445,32 +1433,6 @@
shows "Sup A = fold sup bot A"
using assms sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2)
-lemma Inf_fin_Inf:
- assumes "finite A" and "A \<noteq> {}"
- shows "\<Sqinter>\<^bsub>fin\<^esub>A = Inf A"
-proof -
- interpret ab_semigroup_idem_mult inf
- by (rule ab_semigroup_idem_mult_inf)
- from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
- moreover with `finite A` have "finite B" by simp
- ultimately show ?thesis
- by (simp add: Inf_fin_def fold1_eq_fold_idem inf_Inf_fold_inf [symmetric])
- (simp add: Inf_fold_inf)
-qed
-
-lemma Sup_fin_Sup:
- assumes "finite A" and "A \<noteq> {}"
- shows "\<Squnion>\<^bsub>fin\<^esub>A = Sup A"
-proof -
- interpret ab_semigroup_idem_mult sup
- by (rule ab_semigroup_idem_mult_sup)
- from `A \<noteq> {}` obtain b B where "A = insert b B" by auto
- moreover with `finite A` have "finite B" by simp
- ultimately show ?thesis
- by (simp add: Sup_fin_def fold1_eq_fold_idem sup_Sup_fold_sup [symmetric])
- (simp add: Sup_fold_sup)
-qed
-
lemma inf_INFI_fold_inf:
assumes "finite A"
shows "inf B (INFI A f) = fold (\<lambda>A. inf (f A)) B A" (is "?inf = ?fold")
@@ -3505,4 +1467,661 @@
end
+
+subsection {* Locales as mini-packages *}
+
+locale folding =
+ fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ fixes F :: "'a set \<Rightarrow> 'b \<Rightarrow> 'b"
+ assumes commute_comp: "f x \<circ> f y = f y \<circ> f x"
+ assumes eq_fold: "finite A \<Longrightarrow> F A s = fold f s A"
+begin
+
+lemma fun_left_commute:
+ "f x (f y s) = f y (f x s)"
+ using commute_comp [of x y] by (simp add: expand_fun_eq)
+
+lemma fun_left_comm:
+ "fun_left_comm f"
+proof
+qed (fact fun_left_commute)
+
+lemma empty [simp]:
+ "F {} = id"
+ by (simp add: eq_fold expand_fun_eq)
+
+lemma insert [simp]:
+ assumes "finite A" and "x \<notin> A"
+ shows "F (insert x A) = F A \<circ> f x"
+proof -
+ interpret fun_left_comm f by (fact fun_left_comm)
+ from fold_insert2 assms
+ have "\<And>s. fold f s (insert x A) = fold f (f x s) A" .
+ with `finite A` show ?thesis by (simp add: eq_fold expand_fun_eq)
+qed
+
+lemma remove:
+ assumes "finite A" and "x \<in> A"
+ shows "F A = F (A - {x}) \<circ> f x"
+proof -
+ from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ by (auto dest: mk_disjoint_insert)
+ moreover from `finite A` this have "finite B" by simp
+ ultimately show ?thesis by simp
+qed
+
+lemma insert_remove:
+ assumes "finite A"
+ shows "F (insert x A) = F (A - {x}) \<circ> f x"
+ using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
+
+lemma commute_comp':
+ assumes "finite A"
+ shows "f x \<circ> F A = F A \<circ> f x"
+proof (rule ext)
+ fix s
+ from assms show "(f x \<circ> F A) s = (F A \<circ> f x) s"
+ by (induct A arbitrary: s) (simp_all add: fun_left_commute)
+qed
+
+lemma fun_left_commute':
+ assumes "finite A"
+ shows "f x (F A s) = F A (f x s)"
+ using commute_comp' assms by (simp add: expand_fun_eq)
+
+lemma union:
+ assumes "finite A" and "finite B"
+ and "A \<inter> B = {}"
+ shows "F (A \<union> B) = F A \<circ> F B"
+using `finite A` `A \<inter> B = {}` proof (induct A)
+ case empty show ?case by simp
+next
+ case (insert x A)
+ then have "A \<inter> B = {}" by auto
+ with insert(3) have "F (A \<union> B) = F A \<circ> F B" .
+ moreover from insert have "x \<notin> B" by simp
+ moreover from `finite A` `finite B` have fin: "finite (A \<union> B)" by simp
+ moreover from `x \<notin> A` `x \<notin> B` have "x \<notin> A \<union> B" by simp
+ ultimately show ?case by (simp add: fun_left_commute')
+qed
+
end
+
+locale folding_idem = folding +
+ assumes idem_comp: "f x \<circ> f x = f x"
+begin
+
+declare insert [simp del]
+
+lemma fun_idem:
+ "f x (f x s) = f x s"
+ using idem_comp [of x] by (simp add: expand_fun_eq)
+
+lemma fun_left_comm_idem:
+ "fun_left_comm_idem f"
+proof
+qed (fact fun_left_commute fun_idem)+
+
+lemma insert_idem [simp]:
+ assumes "finite A"
+ shows "F (insert x A) = F A \<circ> f x"
+proof -
+ interpret fun_left_comm_idem f by (fact fun_left_comm_idem)
+ from fold_insert_idem2 assms
+ have "\<And>s. fold f s (insert x A) = fold f (f x s) A" .
+ with assms show ?thesis by (simp add: eq_fold expand_fun_eq)
+qed
+
+lemma union_idem:
+ assumes "finite A" and "finite B"
+ shows "F (A \<union> B) = F A \<circ> F B"
+using `finite A` proof (induct A)
+ case empty show ?case by simp
+next
+ case (insert x A)
+ from insert(3) have "F (A \<union> B) = F A \<circ> F B" .
+ moreover from `finite A` `finite B` have fin: "finite (A \<union> B)" by simp
+ ultimately show ?case by (simp add: fun_left_commute')
+qed
+
+end
+
+no_notation (in times) times (infixl "*" 70)
+no_notation (in one) Groups.one ("1")
+
+locale folding_image_simple = comm_monoid +
+ fixes g :: "('b \<Rightarrow> 'a)"
+ fixes F :: "'b set \<Rightarrow> 'a"
+ assumes eq_fold: "finite A \<Longrightarrow> F A = fold_image f g 1 A"
+begin
+
+lemma empty [simp]:
+ "F {} = 1"
+ by (simp add: eq_fold)
+
+lemma insert [simp]:
+ assumes "finite A" and "x \<notin> A"
+ shows "F (insert x A) = g x * F A"
+proof -
+ interpret fun_left_comm "%x y. (g x) * y" proof
+ qed (simp add: ac_simps)
+ with assms have "fold_image (op *) g 1 (insert x A) = g x * fold_image (op *) g 1 A"
+ by (simp add: fold_image_def)
+ with `finite A` show ?thesis by (simp add: eq_fold)
+qed
+
+lemma remove:
+ assumes "finite A" and "x \<in> A"
+ shows "F A = g x * F (A - {x})"
+proof -
+ from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
+ by (auto dest: mk_disjoint_insert)
+ moreover from `finite A` this have "finite B" by simp
+ ultimately show ?thesis by simp
+qed
+
+lemma insert_remove:
+ assumes "finite A"
+ shows "F (insert x A) = g x * F (A - {x})"
+ using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
+
+lemma union_inter:
+ assumes "finite A" and "finite B"
+ shows "F A * F B = F (A \<union> B) * F (A \<inter> B)"
+using assms proof (induct A)
+ case empty then show ?case by simp
+next
+ case (insert x A) then show ?case
+ by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute)
+qed
+
+corollary union_disjoint:
+ assumes "finite A" and "finite B"
+ assumes "A \<inter> B = {}"
+ shows "F (A \<union> B) = F A * F B"
+ using assms by (simp add: union_inter)
+
+end
+
+locale folding_image = comm_monoid +
+ fixes F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
+ assumes eq_fold: "\<And>g. finite A \<Longrightarrow> F g A = fold_image f g 1 A"
+
+sublocale folding_image < folding_image_simple "op *" 1 g "F g" proof
+qed (fact eq_fold)
+
+context folding_image
+begin
+
+lemma reindex:
+ assumes "finite A" and "inj_on h A"
+ shows "F g (h ` A) = F (g \<circ> h) A"
+ using assms by (induct A) auto
+
+lemma cong:
+ assumes "finite A" and "\<And>x. x \<in> A \<Longrightarrow> g x = h x"
+ shows "F g A = F h A"
+proof -
+ from assms have "ALL C. C <= A --> (ALL x:C. g x = h x) --> F g C = F h C"
+ apply - apply (erule finite_induct) apply simp
+ apply (simp add: subset_insert_iff, clarify)
+ apply (subgoal_tac "finite C")
+ prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl])
+ apply (subgoal_tac "C = insert x (C - {x})")
+ prefer 2 apply blast
+ apply (erule ssubst)
+ apply (drule spec)
+ apply (erule (1) notE impE)
+ apply (simp add: Ball_def del: insert_Diff_single)
+ done
+ with assms show ?thesis by simp
+qed
+
+lemma UNION_disjoint:
+ assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
+ and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
+ shows "F g (UNION I A) = F (F g \<circ> A) I"
+apply (insert assms)
+apply (induct set: finite, simp, atomize)
+apply (subgoal_tac "\<forall>i\<in>Fa. x \<noteq> i")
+ prefer 2 apply blast
+apply (subgoal_tac "A x Int UNION Fa A = {}")
+ prefer 2 apply blast
+apply (simp add: union_disjoint)
+done
+
+lemma distrib:
+ assumes "finite A"
+ shows "F (\<lambda>x. g x * h x) A = F g A * F h A"
+ using assms by (rule finite_induct) (simp_all add: assoc commute left_commute)
+
+lemma related:
+ assumes Re: "R 1 1"
+ and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 * y1) (x2 * y2)"
+ and fS: "finite S" and Rfg: "\<forall>x\<in>S. R (h x) (g x)"
+ shows "R (F h S) (F g S)"
+ using fS by (rule finite_subset_induct) (insert assms, auto)
+
+lemma eq_general:
+ assumes fS: "finite S"
+ and h: "\<forall>y\<in>S'. \<exists>!x. x \<in> S \<and> h x = y"
+ and f12: "\<forall>x\<in>S. h x \<in> S' \<and> f2 (h x) = f1 x"
+ shows "F f1 S = F f2 S'"
+proof-
+ from h f12 have hS: "h ` S = S'" by blast
+ {fix x y assume H: "x \<in> S" "y \<in> S" "h x = h y"
+ from f12 h H have "x = y" by auto }
+ hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast
+ from f12 have th: "\<And>x. x \<in> S \<Longrightarrow> (f2 \<circ> h) x = f1 x" by auto
+ from hS have "F f2 S' = F f2 (h ` S)" by simp
+ also have "\<dots> = F (f2 o h) S" using reindex [OF fS hinj, of f2] .
+ also have "\<dots> = F f1 S " using th cong [OF fS, of "f2 o h" f1]
+ by blast
+ finally show ?thesis ..
+qed
+
+lemma eq_general_inverses:
+ assumes fS: "finite S"
+ and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
+ and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = j x"
+ shows "F j S = F g T"
+ (* metis solves it, but not yet available here *)
+ apply (rule eq_general [OF fS, of T h g j])
+ apply (rule ballI)
+ apply (frule kh)
+ apply (rule ex1I[])
+ apply blast
+ apply clarsimp
+ apply (drule hk) apply simp
+ apply (rule sym)
+ apply (erule conjunct1[OF conjunct2[OF hk]])
+ apply (rule ballI)
+ apply (drule hk)
+ apply blast
+ done
+
+end
+
+notation (in times) times (infixl "*" 70)
+notation (in one) Groups.one ("1")
+
+
+subsection {* Finite cardinality *}
+
+text {* This definition, although traditional, is ugly to work with:
+@{text "card A == LEAST n. EX f. A = {f i | i. i < n}"}.
+But now that we have @{text fold_image} things are easy:
+*}
+
+definition card :: "'a set \<Rightarrow> nat" where
+ "card A = (if finite A then fold_image (op +) (\<lambda>x. 1) 0 A else 0)"
+
+interpretation card!: folding_image_simple "op +" 0 "\<lambda>x. 1" card proof
+qed (simp add: card_def)
+
+lemma card_infinite [simp]:
+ "\<not> finite A \<Longrightarrow> card A = 0"
+ by (simp add: card_def)
+
+lemma card_empty:
+ "card {} = 0"
+ by (fact card.empty)
+
+lemma card_insert_disjoint:
+ "finite A ==> x \<notin> A ==> card (insert x A) = Suc (card A)"
+ by simp
+
+lemma card_insert_if:
+ "finite A ==> card (insert x A) = (if x \<in> A then card A else Suc (card A))"
+ by auto (simp add: card.insert_remove card.remove)
+
+lemma card_ge_0_finite:
+ "card A > 0 \<Longrightarrow> finite A"
+ by (rule ccontr) simp
+
+lemma card_0_eq [simp, noatp]:
+ "finite A \<Longrightarrow> card A = 0 \<longleftrightarrow> A = {}"
+ by (auto dest: mk_disjoint_insert)
+
+lemma finite_UNIV_card_ge_0:
+ "finite (UNIV :: 'a set) \<Longrightarrow> card (UNIV :: 'a set) > 0"
+ by (rule ccontr) simp
+
+lemma card_eq_0_iff:
+ "card A = 0 \<longleftrightarrow> A = {} \<or> \<not> finite A"
+ by auto
+
+lemma card_gt_0_iff:
+ "0 < card A \<longleftrightarrow> A \<noteq> {} \<and> finite A"
+ by (simp add: neq0_conv [symmetric] card_eq_0_iff)
+
+lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
+apply(rule_tac t = A in insert_Diff [THEN subst], assumption)
+apply(simp del:insert_Diff_single)
+done
+
+lemma card_Diff_singleton:
+ "finite A ==> x: A ==> card (A - {x}) = card A - 1"
+by (simp add: card_Suc_Diff1 [symmetric])
+
+lemma card_Diff_singleton_if:
+ "finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
+by (simp add: card_Diff_singleton)
+
+lemma card_Diff_insert[simp]:
+assumes "finite A" and "a:A" and "a ~: B"
+shows "card(A - insert a B) = card(A - B) - 1"
+proof -
+ have "A - insert a B = (A - B) - {a}" using assms by blast
+ then show ?thesis using assms by(simp add:card_Diff_singleton)
+qed
+
+lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
+by (simp add: card_insert_if card_Suc_Diff1 del:card_Diff_insert)
+
+lemma card_insert_le: "finite A ==> card A <= card (insert x A)"
+by (simp add: card_insert_if)
+
+lemma card_mono:
+ assumes "finite B" and "A \<subseteq> B"
+ shows "card A \<le> card B"
+proof -
+ from assms have "finite A" by (auto intro: finite_subset)
+ then show ?thesis using assms proof (induct A arbitrary: B)
+ case empty then show ?case by simp
+ next
+ case (insert x A)
+ then have "x \<in> B" by simp
+ from insert have "A \<subseteq> B - {x}" and "finite (B - {x})" by auto
+ with insert.hyps have "card A \<le> card (B - {x})" by auto
+ with `finite A` `x \<notin> A` `finite B` `x \<in> B` show ?case by simp (simp only: card.remove)
+ qed
+qed
+
+lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)"
+apply (induct set: finite, simp, clarify)
+apply (subgoal_tac "finite A & A - {x} <= F")
+ prefer 2 apply (blast intro: finite_subset, atomize)
+apply (drule_tac x = "A - {x}" in spec)
+apply (simp add: card_Diff_singleton_if split add: split_if_asm)
+apply (case_tac "card A", auto)
+done
+
+lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B"
+apply (simp add: psubset_eq linorder_not_le [symmetric])
+apply (blast dest: card_seteq)
+done
+
+lemma card_Un_Int: "finite A ==> finite B
+ ==> card A + card B = card (A Un B) + card (A Int B)"
+ by (fact card.union_inter)
+
+lemma card_Un_disjoint: "finite A ==> finite B
+ ==> A Int B = {} ==> card (A Un B) = card A + card B"
+ by (fact card.union_disjoint)
+
+lemma card_Diff_subset:
+ assumes "finite B" and "B \<subseteq> A"
+ shows "card (A - B) = card A - card B"
+proof (cases "finite A")
+ case False with assms show ?thesis by simp
+next
+ case True with assms show ?thesis by (induct B arbitrary: A) simp_all
+qed
+
+lemma card_Diff_subset_Int:
+ assumes AB: "finite (A \<inter> B)" shows "card (A - B) = card A - card (A \<inter> B)"
+proof -
+ have "A - B = A - A \<inter> B" by auto
+ thus ?thesis
+ by (simp add: card_Diff_subset AB)
+qed
+
+lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
+apply (rule Suc_less_SucD)
+apply (simp add: card_Suc_Diff1 del:card_Diff_insert)
+done
+
+lemma card_Diff2_less:
+ "finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
+apply (case_tac "x = y")
+ apply (simp add: card_Diff1_less del:card_Diff_insert)
+apply (rule less_trans)
+ prefer 2 apply (auto intro!: card_Diff1_less simp del:card_Diff_insert)
+done
+
+lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
+apply (case_tac "x : A")
+ apply (simp_all add: card_Diff1_less less_imp_le)
+done
+
+lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B"
+by (erule psubsetI, blast)
+
+lemma insert_partition:
+ "\<lbrakk> x \<notin> F; \<forall>c1 \<in> insert x F. \<forall>c2 \<in> insert x F. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {} \<rbrakk>
+ \<Longrightarrow> x \<inter> \<Union> F = {}"
+by auto
+
+lemma finite_psubset_induct[consumes 1, case_names psubset]:
+ assumes "finite A" and "!!A. finite A \<Longrightarrow> (!!B. finite B \<Longrightarrow> B \<subset> A \<Longrightarrow> P(B)) \<Longrightarrow> P(A)" shows "P A"
+using assms(1)
+proof (induct A rule: measure_induct_rule[where f=card])
+ case (less A)
+ show ?case
+ proof(rule assms(2)[OF less(2)])
+ fix B assume "finite B" "B \<subset> A"
+ show "P B" by(rule less(1)[OF psubset_card_mono[OF less(2) `B \<subset> A`] `finite B`])
+ qed
+qed
+
+text{* main cardinality theorem *}
+lemma card_partition [rule_format]:
+ "finite C ==>
+ finite (\<Union> C) -->
+ (\<forall>c\<in>C. card c = k) -->
+ (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = {}) -->
+ k * card(C) = card (\<Union> C)"
+apply (erule finite_induct, simp)
+apply (simp add: card_Un_disjoint insert_partition
+ finite_subset [of _ "\<Union> (insert x F)"])
+done
+
+lemma card_eq_UNIV_imp_eq_UNIV:
+ assumes fin: "finite (UNIV :: 'a set)"
+ and card: "card A = card (UNIV :: 'a set)"
+ shows "A = (UNIV :: 'a set)"
+proof
+ show "A \<subseteq> UNIV" by simp
+ show "UNIV \<subseteq> A"
+ proof
+ fix x
+ show "x \<in> A"
+ proof (rule ccontr)
+ assume "x \<notin> A"
+ then have "A \<subset> UNIV" by auto
+ with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono)
+ with card show False by simp
+ qed
+ qed
+qed
+
+text{*The form of a finite set of given cardinality*}
+
+lemma card_eq_SucD:
+assumes "card A = Suc k"
+shows "\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={})"
+proof -
+ have fin: "finite A" using assms by (auto intro: ccontr)
+ moreover have "card A \<noteq> 0" using assms by auto
+ ultimately obtain b where b: "b \<in> A" by auto
+ show ?thesis
+ proof (intro exI conjI)
+ show "A = insert b (A-{b})" using b by blast
+ show "b \<notin> A - {b}" by blast
+ show "card (A - {b}) = k" and "k = 0 \<longrightarrow> A - {b} = {}"
+ using assms b fin by(fastsimp dest:mk_disjoint_insert)+
+ qed
+qed
+
+lemma card_Suc_eq:
+ "(card A = Suc k) =
+ (\<exists>b B. A = insert b B & b \<notin> B & card B = k & (k=0 \<longrightarrow> B={}))"
+apply(rule iffI)
+ apply(erule card_eq_SucD)
+apply(auto)
+apply(subst card_insert)
+ apply(auto intro:ccontr)
+done
+
+lemma finite_fun_UNIVD2:
+ assumes fin: "finite (UNIV :: ('a \<Rightarrow> 'b) set)"
+ shows "finite (UNIV :: 'b set)"
+proof -
+ from fin have "finite (range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary))"
+ by(rule finite_imageI)
+ moreover have "UNIV = range (\<lambda>f :: 'a \<Rightarrow> 'b. f arbitrary)"
+ by(rule UNIV_eq_I) auto
+ ultimately show "finite (UNIV :: 'b set)" by simp
+qed
+
+lemma card_UNIV_unit: "card (UNIV :: unit set) = 1"
+ unfolding UNIV_unit by simp
+
+
+subsubsection {* Cardinality of image *}
+
+lemma card_image_le: "finite A ==> card (f ` A) <= card A"
+apply (induct set: finite)
+ apply simp
+apply (simp add: le_SucI card_insert_if)
+done
+
+lemma card_image:
+ assumes "inj_on f A"
+ shows "card (f ` A) = card A"
+proof (cases "finite A")
+ case True then show ?thesis using assms by (induct A) simp_all
+next
+ case False then have "\<not> finite (f ` A)" using assms by (auto dest: finite_imageD)
+ with False show ?thesis by simp
+qed
+
+lemma bij_betw_same_card: "bij_betw f A B \<Longrightarrow> card A = card B"
+by(auto simp: card_image bij_betw_def)
+
+lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A"
+by (simp add: card_seteq card_image)
+
+lemma eq_card_imp_inj_on:
+ "[| finite A; card(f ` A) = card A |] ==> inj_on f A"
+apply (induct rule:finite_induct)
+apply simp
+apply(frule card_image_le[where f = f])
+apply(simp add:card_insert_if split:if_splits)
+done
+
+lemma inj_on_iff_eq_card:
+ "finite A ==> inj_on f A = (card(f ` A) = card A)"
+by(blast intro: card_image eq_card_imp_inj_on)
+
+
+lemma card_inj_on_le:
+ "[|inj_on f A; f ` A \<subseteq> B; finite B |] ==> card A \<le> card B"
+apply (subgoal_tac "finite A")
+ apply (force intro: card_mono simp add: card_image [symmetric])
+apply (blast intro: finite_imageD dest: finite_subset)
+done
+
+lemma card_bij_eq:
+ "[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A;
+ finite A; finite B |] ==> card A = card B"
+by (auto intro: le_antisym card_inj_on_le)
+
+
+subsubsection {* Cardinality of sums *}
+
+lemma card_Plus:
+ assumes "finite A" and "finite B"
+ shows "card (A <+> B) = card A + card B"
+proof -
+ have "Inl`A \<inter> Inr`B = {}" by fast
+ with assms show ?thesis
+ unfolding Plus_def
+ by (simp add: card_Un_disjoint card_image)
+qed
+
+lemma card_Plus_conv_if:
+ "card (A <+> B) = (if finite A \<and> finite B then card A + card B else 0)"
+ by (auto simp add: card_Plus)
+
+
+subsubsection {* Cardinality of the Powerset *}
+
+lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *)
+apply (induct set: finite)
+ apply (simp_all add: Pow_insert)
+apply (subst card_Un_disjoint, blast)
+ apply (blast intro: finite_imageI, blast)
+apply (subgoal_tac "inj_on (insert x) (Pow F)")
+ apply (simp add: card_image Pow_insert)
+apply (unfold inj_on_def)
+apply (blast elim!: equalityE)
+done
+
+text {* Relates to equivalence classes. Based on a theorem of F. Kammüller. *}
+
+lemma dvd_partition:
+ "finite (Union C) ==>
+ ALL c : C. k dvd card c ==>
+ (ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
+ k dvd card (Union C)"
+apply(frule finite_UnionD)
+apply(rotate_tac -1)
+apply (induct set: finite, simp_all, clarify)
+apply (subst card_Un_disjoint)
+ apply (auto simp add: disjoint_eq_subset_Compl)
+done
+
+
+subsubsection {* Relating injectivity and surjectivity *}
+
+lemma finite_surj_inj: "finite(A) \<Longrightarrow> A <= f`A \<Longrightarrow> inj_on f A"
+apply(rule eq_card_imp_inj_on, assumption)
+apply(frule finite_imageI)
+apply(drule (1) card_seteq)
+ apply(erule card_image_le)
+apply simp
+done
+
+lemma finite_UNIV_surj_inj: fixes f :: "'a \<Rightarrow> 'a"
+shows "finite(UNIV:: 'a set) \<Longrightarrow> surj f \<Longrightarrow> inj f"
+by (blast intro: finite_surj_inj subset_UNIV dest:surj_range)
+
+lemma finite_UNIV_inj_surj: fixes f :: "'a \<Rightarrow> 'a"
+shows "finite(UNIV:: 'a set) \<Longrightarrow> inj f \<Longrightarrow> surj f"
+by(fastsimp simp:surj_def dest!: endo_inj_surj)
+
+corollary infinite_UNIV_nat[iff]: "~finite(UNIV::nat set)"
+proof
+ assume "finite(UNIV::nat set)"
+ with finite_UNIV_inj_surj[of Suc]
+ show False by simp (blast dest: Suc_neq_Zero surjD)
+qed
+
+(* Often leads to bogus ATP proofs because of reduced type information, hence noatp *)
+lemma infinite_UNIV_char_0[noatp]:
+ "\<not> finite (UNIV::'a::semiring_char_0 set)"
+proof
+ assume "finite (UNIV::'a set)"
+ with subset_UNIV have "finite (range of_nat::'a set)"
+ by (rule finite_subset)
+ moreover have "inj (of_nat::nat \<Rightarrow> 'a)"
+ by (simp add: inj_on_def)
+ ultimately have "finite (UNIV::nat set)"
+ by (rule finite_imageD)
+ then show "False"
+ by simp
+qed
+
+end
--- a/src/HOL/GCD.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/GCD.thy Thu Mar 11 19:06:03 2010 +0100
@@ -1358,10 +1358,10 @@
done
lemma lcm_dvd2_nat: "(n::nat) dvd lcm m n"
- using lcm_dvd1_nat [of n m] by (simp only: lcm_nat_def times.commute gcd_nat.commute)
+ using lcm_dvd1_nat [of n m] by (simp only: lcm_nat_def mult.commute gcd_nat.commute)
lemma lcm_dvd2_int: "(n::int) dvd lcm m n"
- using lcm_dvd1_int [of n m] by (simp only: lcm_int_def lcm_nat_def times.commute gcd_nat.commute)
+ using lcm_dvd1_int [of n m] by (simp only: lcm_int_def lcm_nat_def mult.commute gcd_nat.commute)
lemma dvd_lcm_I1_nat[simp]: "(k::nat) dvd m \<Longrightarrow> k dvd lcm m n"
by(metis lcm_dvd1_nat dvd_trans)
--- a/src/HOL/Groups.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Groups.thy Thu Mar 11 19:06:03 2010 +0100
@@ -67,6 +67,18 @@
end
+locale monoid = semigroup +
+ fixes z :: 'a ("1")
+ assumes left_neutral [simp]: "1 * a = a"
+ assumes right_neutral [simp]: "a * 1 = a"
+
+locale comm_monoid = abel_semigroup +
+ fixes z :: 'a ("1")
+ assumes comm_neutral: "a * 1 = a"
+
+sublocale comm_monoid < monoid proof
+qed (simp_all add: commute comm_neutral)
+
subsection {* Generic operations *}
@@ -129,19 +141,19 @@
class semigroup_add = plus +
assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
-sublocale semigroup_add < plus!: semigroup plus proof
+sublocale semigroup_add < add!: semigroup plus proof
qed (fact add_assoc)
class ab_semigroup_add = semigroup_add +
assumes add_commute [algebra_simps]: "a + b = b + a"
-sublocale ab_semigroup_add < plus!: abel_semigroup plus proof
+sublocale ab_semigroup_add < add!: abel_semigroup plus proof
qed (fact add_commute)
context ab_semigroup_add
begin
-lemmas add_left_commute [algebra_simps] = plus.left_commute
+lemmas add_left_commute [algebra_simps] = add.left_commute
theorems add_ac = add_assoc add_commute add_left_commute
@@ -152,19 +164,19 @@
class semigroup_mult = times +
assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)"
-sublocale semigroup_mult < times!: semigroup times proof
+sublocale semigroup_mult < mult!: semigroup times proof
qed (fact mult_assoc)
class ab_semigroup_mult = semigroup_mult +
assumes mult_commute [algebra_simps]: "a * b = b * a"
-sublocale ab_semigroup_mult < times!: abel_semigroup times proof
+sublocale ab_semigroup_mult < mult!: abel_semigroup times proof
qed (fact mult_commute)
context ab_semigroup_mult
begin
-lemmas mult_left_commute [algebra_simps] = times.left_commute
+lemmas mult_left_commute [algebra_simps] = mult.left_commute
theorems mult_ac = mult_assoc mult_commute mult_left_commute
@@ -173,36 +185,42 @@
theorems mult_ac = mult_assoc mult_commute mult_left_commute
class monoid_add = zero + semigroup_add +
- assumes add_0_left [simp]: "0 + a = a"
- and add_0_right [simp]: "a + 0 = a"
+ assumes add_0_left: "0 + a = a"
+ and add_0_right: "a + 0 = a"
+
+sublocale monoid_add < add!: monoid plus 0 proof
+qed (fact add_0_left add_0_right)+
lemma zero_reorient: "0 = x \<longleftrightarrow> x = 0"
by (rule eq_commute)
class comm_monoid_add = zero + ab_semigroup_add +
assumes add_0: "0 + a = a"
-begin
-subclass monoid_add
- proof qed (insert add_0, simp_all add: add_commute)
+sublocale comm_monoid_add < add!: comm_monoid plus 0 proof
+qed (insert add_0, simp add: ac_simps)
-end
+subclass (in comm_monoid_add) monoid_add proof
+qed (fact add.left_neutral add.right_neutral)+
class monoid_mult = one + semigroup_mult +
- assumes mult_1_left [simp]: "1 * a = a"
- assumes mult_1_right [simp]: "a * 1 = a"
+ assumes mult_1_left: "1 * a = a"
+ and mult_1_right: "a * 1 = a"
+
+sublocale monoid_mult < mult!: monoid times 1 proof
+qed (fact mult_1_left mult_1_right)+
lemma one_reorient: "1 = x \<longleftrightarrow> x = 1"
by (rule eq_commute)
class comm_monoid_mult = one + ab_semigroup_mult +
assumes mult_1: "1 * a = a"
-begin
-subclass monoid_mult
- proof qed (insert mult_1, simp_all add: mult_commute)
+sublocale comm_monoid_mult < mult!: comm_monoid times 1 proof
+qed (insert mult_1, simp add: ac_simps)
-end
+subclass (in comm_monoid_mult) monoid_mult proof
+qed (fact mult.left_neutral mult.right_neutral)+
class cancel_semigroup_add = semigroup_add +
assumes add_left_imp_eq: "a + b = a + c \<Longrightarrow> b = c"
--- a/src/HOL/IsaMakefile Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/IsaMakefile Thu Mar 11 19:06:03 2010 +0100
@@ -247,6 +247,7 @@
MAIN_DEPENDENCIES = $(PLAIN_DEPENDENCIES) \
ATP_Linkup.thy \
+ Big_Operators.thy \
Code_Evaluation.thy \
Code_Numeral.thy \
Divides.thy \
--- a/src/HOL/Lattices.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Lattices.thy Thu Mar 11 19:06:03 2010 +0100
@@ -43,7 +43,7 @@
end
-subsection {* Conrete lattices *}
+subsection {* Concrete lattices *}
notation
less_eq (infix "\<sqsubseteq>" 50) and
--- a/src/HOL/Library/Multiset.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Library/Multiset.thy Thu Mar 11 19:06:03 2010 +0100
@@ -1729,4 +1729,33 @@
"M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P"
by (fact multiset_order.less_asym)
+ML {*
+(* Proof.context -> string -> (typ -> term list) -> typ -> term -> term *)
+fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T]))
+ (Const _ $ t') =
+ let
+ val (maybe_opt, ps) =
+ Nitpick_Model.dest_plain_fun t' ||> op ~~
+ ||> map (apsnd (snd o HOLogic.dest_number))
+ fun elems_for t =
+ case AList.lookup (op =) ps t of
+ SOME n => replicate n t
+ | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]
+ in
+ case maps elems_for (all_values elem_T) @
+ (if maybe_opt then [Const (Nitpick_Model.unrep, elem_T)] else []) of
+ [] => Const (@{const_name zero_class.zero}, T)
+ | ts => foldl1 (fn (t1, t2) =>
+ Const (@{const_name plus_class.plus}, T --> T --> T)
+ $ t1 $ t2)
+ (map (curry (op $) (Const (@{const_name single},
+ elem_T --> T))) ts)
+ end
+ | multiset_postproc _ _ _ _ t = t
+*}
+
+setup {*
+Nitpick.register_term_postprocessor @{typ "'a multiset"} multiset_postproc
+*}
+
end
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Thu Mar 11 19:06:03 2010 +0100
@@ -94,8 +94,8 @@
using lem1[unfolded lem3 lem2 lem5] by auto
have even_minus_odd:"\<And>x y. even x \<Longrightarrow> odd (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
have odd_minus_even:"\<And>x y. odd x \<Longrightarrow> even (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
- show ?thesis unfolding even_nat_def unfolding card_def and lem4[THEN sym] and *[unfolded card_def]
- unfolding card_def[THEN sym] apply(rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)
+ show ?thesis unfolding even_nat_def unfolding card_eq_setsum and lem4[THEN sym] and *[unfolded card_eq_setsum]
+ unfolding card_eq_setsum[THEN sym] apply (rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)
apply(rule assms(7)[unfolded even_nat_def]) unfolding int_mult by auto qed
subsection {* The odd/even result for faces of complete vertices, generalized. *}
--- a/src/HOL/Multivariate_Analysis/Integration.cert Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.cert Thu Mar 11 19:06:03 2010 +0100
@@ -3294,3 +3294,215 @@
else -> val!7
}
sat
+qkVrmXMcHAG5MLuJ9d8jXQ 211 0
+#2 := false
+#33 := 0::real
+decl uf_11 :: (-> T5 T6 real)
+decl uf_15 :: T6
+#28 := uf_15
+decl uf_16 :: T5
+#30 := uf_16
+#31 := (uf_11 uf_16 uf_15)
+decl uf_12 :: (-> T7 T8 T5)
+decl uf_14 :: T8
+#26 := uf_14
+decl uf_13 :: (-> T1 T7)
+decl uf_8 :: T1
+#16 := uf_8
+#25 := (uf_13 uf_8)
+#27 := (uf_12 #25 uf_14)
+#29 := (uf_11 #27 uf_15)
+#73 := -1::real
+#84 := (* -1::real #29)
+#85 := (+ #84 #31)
+#74 := (* -1::real #31)
+#75 := (+ #29 #74)
+#112 := (>= #75 0::real)
+#119 := (ite #112 #75 #85)
+#127 := (* -1::real #119)
+decl uf_17 :: T5
+#37 := uf_17
+#38 := (uf_11 uf_17 uf_15)
+#102 := -1/3::real
+#103 := (* -1/3::real #38)
+#128 := (+ #103 #127)
+#100 := 1/3::real
+#101 := (* 1/3::real #31)
+#129 := (+ #101 #128)
+#130 := (<= #129 0::real)
+#131 := (not #130)
+#40 := 3::real
+#39 := (- #31 #38)
+#41 := (/ #39 3::real)
+#32 := (- #29 #31)
+#35 := (- #32)
+#34 := (< #32 0::real)
+#36 := (ite #34 #35 #32)
+#42 := (< #36 #41)
+#136 := (iff #42 #131)
+#104 := (+ #101 #103)
+#78 := (< #75 0::real)
+#90 := (ite #78 #85 #75)
+#109 := (< #90 #104)
+#134 := (iff #109 #131)
+#124 := (< #119 #104)
+#132 := (iff #124 #131)
+#133 := [rewrite]: #132
+#125 := (iff #109 #124)
+#122 := (= #90 #119)
+#113 := (not #112)
+#116 := (ite #113 #85 #75)
+#120 := (= #116 #119)
+#121 := [rewrite]: #120
+#117 := (= #90 #116)
+#114 := (iff #78 #113)
+#115 := [rewrite]: #114
+#118 := [monotonicity #115]: #117
+#123 := [trans #118 #121]: #122
+#126 := [monotonicity #123]: #125
+#135 := [trans #126 #133]: #134
+#110 := (iff #42 #109)
+#107 := (= #41 #104)
+#93 := (* -1::real #38)
+#94 := (+ #31 #93)
+#97 := (/ #94 3::real)
+#105 := (= #97 #104)
+#106 := [rewrite]: #105
+#98 := (= #41 #97)
+#95 := (= #39 #94)
+#96 := [rewrite]: #95
+#99 := [monotonicity #96]: #98
+#108 := [trans #99 #106]: #107
+#91 := (= #36 #90)
+#76 := (= #32 #75)
+#77 := [rewrite]: #76
+#88 := (= #35 #85)
+#81 := (- #75)
+#86 := (= #81 #85)
+#87 := [rewrite]: #86
+#82 := (= #35 #81)
+#83 := [monotonicity #77]: #82
+#89 := [trans #83 #87]: #88
+#79 := (iff #34 #78)
+#80 := [monotonicity #77]: #79
+#92 := [monotonicity #80 #89 #77]: #91
+#111 := [monotonicity #92 #108]: #110
+#137 := [trans #111 #135]: #136
+#72 := [asserted]: #42
+#138 := [mp #72 #137]: #131
+decl uf_1 :: T1
+#4 := uf_1
+#43 := (uf_13 uf_1)
+#44 := (uf_12 #43 uf_14)
+#45 := (uf_11 #44 uf_15)
+#149 := (* -1::real #45)
+#150 := (+ #38 #149)
+#140 := (+ #93 #45)
+#161 := (<= #150 0::real)
+#168 := (ite #161 #140 #150)
+#176 := (* -1::real #168)
+#177 := (+ #103 #176)
+#178 := (+ #101 #177)
+#179 := (<= #178 0::real)
+#180 := (not #179)
+#46 := (- #45 #38)
+#48 := (- #46)
+#47 := (< #46 0::real)
+#49 := (ite #47 #48 #46)
+#50 := (< #49 #41)
+#185 := (iff #50 #180)
+#143 := (< #140 0::real)
+#155 := (ite #143 #150 #140)
+#158 := (< #155 #104)
+#183 := (iff #158 #180)
+#173 := (< #168 #104)
+#181 := (iff #173 #180)
+#182 := [rewrite]: #181
+#174 := (iff #158 #173)
+#171 := (= #155 #168)
+#162 := (not #161)
+#165 := (ite #162 #150 #140)
+#169 := (= #165 #168)
+#170 := [rewrite]: #169
+#166 := (= #155 #165)
+#163 := (iff #143 #162)
+#164 := [rewrite]: #163
+#167 := [monotonicity #164]: #166
+#172 := [trans #167 #170]: #171
+#175 := [monotonicity #172]: #174
+#184 := [trans #175 #182]: #183
+#159 := (iff #50 #158)
+#156 := (= #49 #155)
+#141 := (= #46 #140)
+#142 := [rewrite]: #141
+#153 := (= #48 #150)
+#146 := (- #140)
+#151 := (= #146 #150)
+#152 := [rewrite]: #151
+#147 := (= #48 #146)
+#148 := [monotonicity #142]: #147
+#154 := [trans #148 #152]: #153
+#144 := (iff #47 #143)
+#145 := [monotonicity #142]: #144
+#157 := [monotonicity #145 #154 #142]: #156
+#160 := [monotonicity #157 #108]: #159
+#186 := [trans #160 #184]: #185
+#139 := [asserted]: #50
+#187 := [mp #139 #186]: #180
+#299 := (+ #140 #176)
+#300 := (<= #299 0::real)
+#290 := (= #140 #168)
+#329 := [hypothesis]: #162
+#191 := (+ #29 #149)
+#192 := (<= #191 0::real)
+#51 := (<= #29 #45)
+#193 := (iff #51 #192)
+#194 := [rewrite]: #193
+#188 := [asserted]: #51
+#195 := [mp #188 #194]: #192
+#298 := (+ #75 #127)
+#301 := (<= #298 0::real)
+#284 := (= #75 #119)
+#302 := [hypothesis]: #113
+#296 := (+ #85 #127)
+#297 := (<= #296 0::real)
+#285 := (= #85 #119)
+#288 := (or #112 #285)
+#289 := [def-axiom]: #288
+#303 := [unit-resolution #289 #302]: #285
+#304 := (not #285)
+#305 := (or #304 #297)
+#306 := [th-lemma]: #305
+#307 := [unit-resolution #306 #303]: #297
+#315 := (not #290)
+#310 := (not #300)
+#311 := (or #310 #112)
+#308 := [hypothesis]: #300
+#309 := [th-lemma #308 #307 #138 #302 #187 #195]: false
+#312 := [lemma #309]: #311
+#322 := [unit-resolution #312 #302]: #310
+#316 := (or #315 #300)
+#313 := [hypothesis]: #310
+#314 := [hypothesis]: #290
+#317 := [th-lemma]: #316
+#318 := [unit-resolution #317 #314 #313]: false
+#319 := [lemma #318]: #316
+#323 := [unit-resolution #319 #322]: #315
+#292 := (or #162 #290)
+#293 := [def-axiom]: #292
+#324 := [unit-resolution #293 #323]: #162
+#325 := [th-lemma #324 #307 #138 #302 #195]: false
+#326 := [lemma #325]: #112
+#286 := (or #113 #284)
+#287 := [def-axiom]: #286
+#330 := [unit-resolution #287 #326]: #284
+#331 := (not #284)
+#332 := (or #331 #301)
+#333 := [th-lemma]: #332
+#334 := [unit-resolution #333 #330]: #301
+#335 := [th-lemma #326 #334 #195 #329 #138]: false
+#336 := [lemma #335]: #161
+#327 := [unit-resolution #293 #336]: #290
+#328 := [unit-resolution #319 #327]: #300
+[th-lemma #326 #334 #195 #328 #187 #138]: false
+unsat
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Thu Mar 11 19:06:03 2010 +0100
@@ -118,10 +118,11 @@
oops
ML {*
-(* Proof.context -> typ -> term -> term *)
-fun my_int_postproc _ T (Const _ $ (Const _ $ t1 $ t2)) =
- HOLogic.mk_number T (snd (HOLogic.dest_number t1) - snd (HOLogic.dest_number t2))
- | my_int_postproc _ _ t = t
+(* Proof.context -> string -> (typ -> term list) -> typ -> term -> term *)
+fun my_int_postproc _ _ _ T (Const _ $ (Const _ $ t1 $ t2)) =
+ HOLogic.mk_number T (snd (HOLogic.dest_number t1)
+ - snd (HOLogic.dest_number t2))
+ | my_int_postproc _ _ _ _ t = t
*}
setup {* Nitpick.register_term_postprocessor @{typ my_int} my_int_postproc *}
@@ -279,19 +280,19 @@
"n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
lemma "n \<in> reach \<Longrightarrow> 2 dvd n"
-nitpick [unary_ints, expect = none]
+nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (induct set: reach)
apply auto
- nitpick [expect = none]
+ nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (thin_tac "n \<in> reach")
nitpick [expect = genuine]
oops
lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0"
-nitpick [unary_ints, expect = none]
+nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (induct set: reach)
apply auto
- nitpick [expect = none]
+ nitpick [card = 1\<midarrow>5, bits = 1\<midarrow>5, expect = none]
apply (thin_tac "n \<in> reach")
nitpick [expect = genuine]
oops
--- a/src/HOL/Number_Theory/Binomial.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Number_Theory/Binomial.thy Thu Mar 11 19:06:03 2010 +0100
@@ -364,7 +364,7 @@
finally have "card ({T. T \<le> insert x F \<and> card T = k + 1}) =
card F choose (k + 1) + (card F choose k)".
with iassms choose_plus_one_nat show ?thesis
- by auto
+ by (auto simp del: card.insert)
qed
qed
qed
--- a/src/HOL/Option.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Option.thy Thu Mar 11 19:06:03 2010 +0100
@@ -5,7 +5,7 @@
header {* Datatype option *}
theory Option
-imports Datatype Finite_Set
+imports Datatype
begin
datatype 'a option = None | Some 'a
@@ -33,13 +33,6 @@
lemma UNIV_option_conv: "UNIV = insert None (range Some)"
by(auto intro: classical)
-lemma finite_option_UNIV[simp]:
- "finite (UNIV :: 'a option set) = finite (UNIV :: 'a set)"
-by(auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some)
-
-instance option :: (finite) finite proof
-qed (simp add: UNIV_option_conv)
-
subsubsection {* Operations *}
--- a/src/HOL/Rat.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Rat.thy Thu Mar 11 19:06:03 2010 +0100
@@ -1104,11 +1104,11 @@
lemma rat_less_eq_code [code]:
"p \<le> q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d \<le> c * b)"
- by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+ by (cases p, cases q) (simp add: quotient_of_Fract mult.commute)
lemma rat_less_code [code]:
"p < q \<longleftrightarrow> (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d < c * b)"
- by (cases p, cases q) (simp add: quotient_of_Fract times.commute)
+ by (cases p, cases q) (simp add: quotient_of_Fract mult.commute)
lemma [code]:
"of_rat p = (let (a, b) = quotient_of p in of_int a / of_int b)"
--- a/src/HOL/Tools/Nitpick/HISTORY Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Tools/Nitpick/HISTORY Thu Mar 11 19:06:03 2010 +0100
@@ -7,6 +7,7 @@
* Added support for quotient types
* Added support for local definitions (for "function" and "termination"
proofs)
+ * Added support for term postprocessors
* Optimized "Multiset.multiset" and "FinFun.finfun"
* Improved efficiency of "destroy_constrs" optimization
* Fixed soundness bugs related to "destroy_constrs" optimization and record
--- a/src/HOL/Tools/Nitpick/kodkod.ML Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Thu Mar 11 19:06:03 2010 +0100
@@ -319,6 +319,8 @@
int_expr_func: int_expr -> 'a -> 'a
}
+(** Auxiliary functions on ML representation of Kodkod problems **)
+
(* 'a fold_expr_funcs -> formula -> 'a -> 'a *)
fun fold_formula (F : 'a fold_expr_funcs) formula =
case formula of
@@ -512,18 +514,7 @@
filter (is_relevant_setting o fst) (#settings p1)
= filter (is_relevant_setting o fst) (#settings p2)
-val created_temp_dir = Unsynchronized.ref false
-
-(* bool -> string * string *)
-fun serial_string_and_temporary_dir_for_overlord overlord =
- if overlord then
- ("", getenv "ISABELLE_HOME_USER")
- else
- let
- val dir = getenv "ISABELLE_TMP"
- val _ = if !created_temp_dir then ()
- else (created_temp_dir := true; File.mkdir (Path.explode dir))
- in (serial_string (), dir) end
+(** Serialization of problem **)
(* int -> string *)
fun base_name j =
@@ -901,6 +892,8 @@
map out_problem problems
end
+(** Parsing of solution **)
+
(* string -> bool *)
fun is_ident_char s =
Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse
@@ -1016,6 +1009,21 @@
is partly due to the JVM and partly due to the ML "bash" function. *)
val fudge_ms = 250
+(** Main Kodkod entry point **)
+
+val created_temp_dir = Unsynchronized.ref false
+
+(* bool -> string * string *)
+fun serial_string_and_temporary_dir_for_overlord overlord =
+ if overlord then
+ ("", getenv "ISABELLE_HOME_USER")
+ else
+ let
+ val dir = getenv "ISABELLE_TMP"
+ val _ = if !created_temp_dir then ()
+ else (created_temp_dir := true; File.mkdir (Path.explode dir))
+ in (serial_string (), dir) end
+
(* bool -> Time.time option -> int -> int -> problem list -> outcome *)
fun solve_any_problem overlord deadline max_threads max_solutions problems =
let
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Thu Mar 11 19:06:03 2010 +0100
@@ -52,6 +52,9 @@
val name_sep : string
val numeral_prefix : string
+ val base_prefix : string
+ val step_prefix : string
+ val unrolled_prefix : string
val ubfp_prefix : string
val lbfp_prefix : string
val quot_normal_prefix : string
@@ -59,6 +62,8 @@
val special_prefix : string
val uncurry_prefix : string
val eval_prefix : string
+ val iter_var_prefix : string
+ val strip_first_name_sep : string -> string * string
val original_name : string -> string
val s_conj : term * term -> term
val s_disj : term * term -> term
@@ -169,6 +174,7 @@
val term_under_def : term -> term
val case_const_names :
theory -> (typ option * bool) list -> (string * int) list
+ val unfold_defs_in_term : hol_context -> term -> term
val const_def_table :
Proof.context -> (term * term) list -> term list -> const_table
val const_nondef_table : term list -> const_table
@@ -184,13 +190,13 @@
val optimized_quot_type_axioms :
Proof.context -> (typ option * bool) list -> string * typ list -> term list
val def_of_const : theory -> const_table -> styp -> term option
+ val fixpoint_kind_of_rhs : term -> fixpoint_kind
val fixpoint_kind_of_const :
theory -> const_table -> string * typ -> fixpoint_kind
val is_inductive_pred : hol_context -> styp -> bool
val is_equational_fun : hol_context -> styp -> bool
val is_constr_pattern_lhs : theory -> term -> bool
val is_constr_pattern_formula : theory -> term -> bool
- val unfold_defs_in_term : hol_context -> term -> term
val codatatype_bisim_axioms : hol_context -> typ -> term list
val is_well_founded_inductive_pred : hol_context -> styp -> bool
val unrolled_inductive_pred_const : hol_context -> bool -> styp -> term
@@ -199,13 +205,6 @@
val merge_type_vars_in_terms : term list -> term list
val ground_types_in_type : hol_context -> bool -> typ -> typ list
val ground_types_in_terms : hol_context -> bool -> term list -> typ list
- val format_type : int list -> int list -> typ -> typ
- val format_term_type :
- theory -> const_table -> (term option * int list) list -> term -> typ
- val user_friendly_const :
- hol_context -> string * string -> (term option * int list) list
- -> styp -> term * typ
- val assign_operator_for_const : styp -> string
end;
structure Nitpick_HOL : NITPICK_HOL =
@@ -283,7 +282,8 @@
val uncurry_prefix = nitpick_prefix ^ "unc"
val eval_prefix = nitpick_prefix ^ "eval"
val iter_var_prefix = "i"
-val arg_var_prefix = "x"
+
+(** Constant/type information and term/type manipulation **)
(* int -> string *)
fun sel_prefix_for j = sel_prefix ^ string_of_int j ^ name_sep
@@ -301,7 +301,13 @@
case strip_first_name_sep s of (s1, "") => s1 | (_, s2) => original_name s2
else
s
-val after_name_sep = snd o strip_first_name_sep
+
+(* term * term -> term *)
+fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
+ | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t
+ | s_betapply p = betapply p
+(* term * term list -> term *)
+val s_betapplys = Library.foldl s_betapply
(* term * term -> term *)
fun s_conj (t1, @{const True}) = t1
@@ -505,7 +511,8 @@
Type ((if gfp then gfp_iterator_prefix else lfp_iterator_prefix) ^ s,
binder_types T)
(* typ -> styp *)
-fun const_for_iterator_type (Type (s, Ts)) = (after_name_sep s, Ts ---> bool_T)
+fun const_for_iterator_type (Type (s, Ts)) =
+ (strip_first_name_sep s |> snd, Ts ---> bool_T)
| const_for_iterator_type T =
raise TYPE ("Nitpick_HOL.const_for_iterator_type", [T], [])
@@ -540,13 +547,6 @@
fun dest_n_tuple 1 t = [t]
| dest_n_tuple n t = HOLogic.dest_prod t ||> dest_n_tuple (n - 1) |> op ::
-(* int -> typ -> typ list *)
-fun dest_n_tuple_type 1 T = [T]
- | dest_n_tuple_type n (Type (_, [T1, T2])) =
- T1 :: dest_n_tuple_type (n - 1) T2
- | dest_n_tuple_type _ T =
- raise TYPE ("Nitpick_HOL.dest_n_tuple_type", [T], [])
-
type typedef_info =
{rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string,
set_def: thm option, prop_of_Rep: thm, set_name: string,
@@ -1463,106 +1463,57 @@
(Datatype.get_all thy) [] @
map (apsnd length o snd) (#codatatypes (Data.get thy))
-(* Proof.context -> (term * term) list -> term list -> const_table *)
-fun const_def_table ctxt subst ts =
- table_for (map prop_of o Nitpick_Defs.get) ctxt subst
- |> fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
- (map pair_for_prop ts)
-(* term list -> const_table *)
-fun const_nondef_table ts =
- fold (fn t => append (map (fn s => (s, t)) (Term.add_const_names t []))) ts []
- |> AList.group (op =) |> Symtab.make
-(* Proof.context -> (term * term) list -> const_table *)
-val const_simp_table = table_for (map prop_of o Nitpick_Simps.get)
-val const_psimp_table = table_for (map prop_of o Nitpick_Psimps.get)
-(* Proof.context -> (term * term) list -> const_table -> const_table *)
-fun inductive_intro_table ctxt subst def_table =
- table_for (map (unfold_mutually_inductive_preds (ProofContext.theory_of ctxt)
- def_table o prop_of)
- o Nitpick_Intros.get) ctxt subst
-(* theory -> term list Inttab.table *)
-fun ground_theorem_table thy =
- fold ((fn @{const Trueprop} $ t1 =>
- is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1)
- | _ => I) o prop_of o snd) (PureThy.all_thms_of thy) Inttab.empty
+(* theory -> const_table -> string * typ -> fixpoint_kind *)
+fun fixpoint_kind_of_const thy table x =
+ if is_built_in_const thy [(NONE, false)] false x then
+ NoFp
+ else
+ fixpoint_kind_of_rhs (the (def_of_const thy table x))
+ handle Option.Option => NoFp
-val basic_ersatz_table =
- [(@{const_name prod_case}, @{const_name split}),
- (@{const_name card}, @{const_name card'}),
- (@{const_name setsum}, @{const_name setsum'}),
- (@{const_name fold_graph}, @{const_name fold_graph'}),
- (@{const_name wf}, @{const_name wf'}),
- (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
- (@{const_name wfrec}, @{const_name wfrec'})]
-
-(* theory -> (string * string) list *)
-fun ersatz_table thy =
- fold (append o snd) (#frac_types (Data.get thy)) basic_ersatz_table
-
-(* const_table Unsynchronized.ref -> string -> term list -> unit *)
-fun add_simps simp_table s eqs =
- Unsynchronized.change simp_table
- (Symtab.update (s, eqs @ these (Symtab.lookup (!simp_table) s)))
+(* hol_context -> styp -> bool *)
+fun is_real_inductive_pred ({thy, stds, fast_descrs, def_table, intro_table,
+ ...} : hol_context) x =
+ fixpoint_kind_of_const thy def_table x <> NoFp andalso
+ not (null (def_props_for_const thy stds fast_descrs intro_table x))
+fun is_real_equational_fun ({thy, stds, fast_descrs, simp_table, psimp_table,
+ ...} : hol_context) x =
+ exists (fn table => not (null (def_props_for_const thy stds fast_descrs table
+ x)))
+ [!simp_table, psimp_table]
+fun is_inductive_pred hol_ctxt =
+ is_real_inductive_pred hol_ctxt andf (not o is_real_equational_fun hol_ctxt)
+fun is_equational_fun hol_ctxt =
+ (is_real_equational_fun hol_ctxt orf is_real_inductive_pred hol_ctxt orf
+ (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
-(* theory -> styp -> term list *)
-fun inverse_axioms_for_rep_fun thy (x as (_, T)) =
- let val abs_T = domain_type T in
- typedef_info thy (fst (dest_Type abs_T)) |> the
- |> pairf #Abs_inverse #Rep_inverse
- |> pairself (Refute.specialize_type thy x o prop_of o the)
- ||> single |> op ::
- end
-(* theory -> string * typ list -> term list *)
-fun optimized_typedef_axioms thy (abs_z as (abs_s, _)) =
- let val abs_T = Type abs_z in
- if is_univ_typedef thy abs_T then
- []
- else case typedef_info thy abs_s of
- SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
- let
- val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
- val rep_t = Const (Rep_name, abs_T --> rep_T)
- val set_t = Const (set_name, rep_T --> bool_T)
- val set_t' =
- prop_of_Rep |> HOLogic.dest_Trueprop
- |> Refute.specialize_type thy (dest_Const rep_t)
- |> HOLogic.dest_mem |> snd
- in
- [HOLogic.all_const abs_T
- $ Abs (Name.uu, abs_T, set_t $ (rep_t $ Bound 0))]
- |> set_t <> set_t' ? cons (HOLogic.mk_eq (set_t, set_t'))
- |> map HOLogic.mk_Trueprop
- end
- | NONE => []
- end
-(* Proof.context -> string * typ list -> term list *)
-fun optimized_quot_type_axioms ctxt stds abs_z =
- let
- val thy = ProofContext.theory_of ctxt
- val abs_T = Type abs_z
- val rep_T = rep_type_for_quot_type thy abs_T
- val equiv_rel = equiv_relation_for_quot_type thy abs_T
- val a_var = Var (("a", 0), abs_T)
- val x_var = Var (("x", 0), rep_T)
- val y_var = Var (("y", 0), rep_T)
- val x = (@{const_name Quot}, rep_T --> abs_T)
- val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
- val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
- val normal_x = normal_t $ x_var
- val normal_y = normal_t $ y_var
- val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
- in
- [Logic.mk_equals (normal_t $ sel_a_t, sel_a_t),
- Logic.list_implies
- ([@{const Not} $ (is_unknown_t $ normal_x),
- @{const Not} $ (is_unknown_t $ normal_y),
- equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
- Logic.mk_equals (normal_x, normal_y)),
- Logic.list_implies
- ([HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)),
- HOLogic.mk_Trueprop (@{const Not} $ HOLogic.mk_eq (normal_x, x_var))],
- HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
- end
+(* term -> term *)
+fun lhs_of_equation t =
+ case t of
+ Const (@{const_name all}, _) $ Abs (_, _, t1) => lhs_of_equation t1
+ | Const (@{const_name "=="}, _) $ t1 $ _ => SOME t1
+ | @{const "==>"} $ _ $ t2 => lhs_of_equation t2
+ | @{const Trueprop} $ t1 => lhs_of_equation t1
+ | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
+ | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
+ | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
+ | _ => NONE
+(* theory -> term -> bool *)
+fun is_constr_pattern _ (Bound _) = true
+ | is_constr_pattern _ (Var _) = true
+ | is_constr_pattern thy t =
+ case strip_comb t of
+ (Const x, args) =>
+ is_constr_like thy x andalso forall (is_constr_pattern thy) args
+ | _ => false
+fun is_constr_pattern_lhs thy t =
+ forall (is_constr_pattern thy) (snd (strip_comb t))
+fun is_constr_pattern_formula thy t =
+ case lhs_of_equation t of
+ SOME t' => is_constr_pattern_lhs thy t'
+ | NONE => false
+
+(** Constant unfolding **)
(* theory -> (typ option * bool) list -> int * styp -> term *)
fun constr_case_body thy stds (j, (x as (_, T))) =
@@ -1586,7 +1537,6 @@
|> fold_rev (add_constr_case hol_ctxt res_T) (length xs downto 2 ~~ xs')
|> fold_rev (curry absdummy) (func_Ts @ [dataT])
end
-
(* hol_context -> string -> typ -> typ -> term -> term *)
fun optimized_record_get (hol_ctxt as {thy, stds, ...}) s rec_T res_T t =
let val constr_x = hd (datatype_constrs hol_ctxt rec_T) in
@@ -1624,63 +1574,6 @@
end) (index_seq 0 n) Ts
in list_comb (Const constr_x, ts) end
-(* theory -> const_table -> string * typ -> fixpoint_kind *)
-fun fixpoint_kind_of_const thy table x =
- if is_built_in_const thy [(NONE, false)] false x then
- NoFp
- else
- fixpoint_kind_of_rhs (the (def_of_const thy table x))
- handle Option.Option => NoFp
-
-(* hol_context -> styp -> bool *)
-fun is_real_inductive_pred ({thy, stds, fast_descrs, def_table, intro_table,
- ...} : hol_context) x =
- fixpoint_kind_of_const thy def_table x <> NoFp andalso
- not (null (def_props_for_const thy stds fast_descrs intro_table x))
-fun is_real_equational_fun ({thy, stds, fast_descrs, simp_table, psimp_table,
- ...} : hol_context) x =
- exists (fn table => not (null (def_props_for_const thy stds fast_descrs table
- x)))
- [!simp_table, psimp_table]
-fun is_inductive_pred hol_ctxt =
- is_real_inductive_pred hol_ctxt andf (not o is_real_equational_fun hol_ctxt)
-fun is_equational_fun hol_ctxt =
- (is_real_equational_fun hol_ctxt orf is_real_inductive_pred hol_ctxt orf
- (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
-
-(* term * term -> term *)
-fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
- | s_betapply (Const (@{const_name If}, _) $ @{const False} $ _, t) = t
- | s_betapply p = betapply p
-(* term * term list -> term *)
-val s_betapplys = Library.foldl s_betapply
-
-(* term -> term *)
-fun lhs_of_equation t =
- case t of
- Const (@{const_name all}, _) $ Abs (_, _, t1) => lhs_of_equation t1
- | Const (@{const_name "=="}, _) $ t1 $ _ => SOME t1
- | @{const "==>"} $ _ $ t2 => lhs_of_equation t2
- | @{const Trueprop} $ t1 => lhs_of_equation t1
- | Const (@{const_name All}, _) $ Abs (_, _, t1) => lhs_of_equation t1
- | Const (@{const_name "op ="}, _) $ t1 $ _ => SOME t1
- | @{const "op -->"} $ _ $ t2 => lhs_of_equation t2
- | _ => NONE
-(* theory -> term -> bool *)
-fun is_constr_pattern _ (Bound _) = true
- | is_constr_pattern _ (Var _) = true
- | is_constr_pattern thy t =
- case strip_comb t of
- (Const x, args) =>
- is_constr_like thy x andalso forall (is_constr_pattern thy) args
- | _ => false
-fun is_constr_pattern_lhs thy t =
- forall (is_constr_pattern thy) (snd (strip_comb t))
-fun is_constr_pattern_formula thy t =
- case lhs_of_equation t of
- SOME t' => is_constr_pattern_lhs thy t'
- | NONE => false
-
(* Prevents divergence in case of cyclic or infinite definition dependencies. *)
val unfold_max_depth = 255
@@ -1823,6 +1716,109 @@
in s_betapplys (const, map (do_term depth Ts) ts) |> Envir.beta_norm end
in do_term 0 [] end
+(** Axiom extraction/generation **)
+
+(* Proof.context -> (term * term) list -> term list -> const_table *)
+fun const_def_table ctxt subst ts =
+ table_for (map prop_of o Nitpick_Defs.get) ctxt subst
+ |> fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
+ (map pair_for_prop ts)
+(* term list -> const_table *)
+fun const_nondef_table ts =
+ fold (fn t => append (map (fn s => (s, t)) (Term.add_const_names t []))) ts []
+ |> AList.group (op =) |> Symtab.make
+(* Proof.context -> (term * term) list -> const_table *)
+val const_simp_table = table_for (map prop_of o Nitpick_Simps.get)
+val const_psimp_table = table_for (map prop_of o Nitpick_Psimps.get)
+(* Proof.context -> (term * term) list -> const_table -> const_table *)
+fun inductive_intro_table ctxt subst def_table =
+ table_for (map (unfold_mutually_inductive_preds (ProofContext.theory_of ctxt)
+ def_table o prop_of)
+ o Nitpick_Intros.get) ctxt subst
+(* theory -> term list Inttab.table *)
+fun ground_theorem_table thy =
+ fold ((fn @{const Trueprop} $ t1 =>
+ is_ground_term t1 ? Inttab.map_default (hash_term t1, []) (cons t1)
+ | _ => I) o prop_of o snd) (PureThy.all_thms_of thy) Inttab.empty
+
+val basic_ersatz_table =
+ [(@{const_name prod_case}, @{const_name split}),
+ (@{const_name card}, @{const_name card'}),
+ (@{const_name setsum}, @{const_name setsum'}),
+ (@{const_name fold_graph}, @{const_name fold_graph'}),
+ (@{const_name wf}, @{const_name wf'}),
+ (@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
+ (@{const_name wfrec}, @{const_name wfrec'})]
+
+(* theory -> (string * string) list *)
+fun ersatz_table thy =
+ fold (append o snd) (#frac_types (Data.get thy)) basic_ersatz_table
+
+(* const_table Unsynchronized.ref -> string -> term list -> unit *)
+fun add_simps simp_table s eqs =
+ Unsynchronized.change simp_table
+ (Symtab.update (s, eqs @ these (Symtab.lookup (!simp_table) s)))
+
+(* theory -> styp -> term list *)
+fun inverse_axioms_for_rep_fun thy (x as (_, T)) =
+ let val abs_T = domain_type T in
+ typedef_info thy (fst (dest_Type abs_T)) |> the
+ |> pairf #Abs_inverse #Rep_inverse
+ |> pairself (Refute.specialize_type thy x o prop_of o the)
+ ||> single |> op ::
+ end
+(* theory -> string * typ list -> term list *)
+fun optimized_typedef_axioms thy (abs_z as (abs_s, _)) =
+ let val abs_T = Type abs_z in
+ if is_univ_typedef thy abs_T then
+ []
+ else case typedef_info thy abs_s of
+ SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
+ let
+ val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
+ val rep_t = Const (Rep_name, abs_T --> rep_T)
+ val set_t = Const (set_name, rep_T --> bool_T)
+ val set_t' =
+ prop_of_Rep |> HOLogic.dest_Trueprop
+ |> Refute.specialize_type thy (dest_Const rep_t)
+ |> HOLogic.dest_mem |> snd
+ in
+ [HOLogic.all_const abs_T
+ $ Abs (Name.uu, abs_T, set_t $ (rep_t $ Bound 0))]
+ |> set_t <> set_t' ? cons (HOLogic.mk_eq (set_t, set_t'))
+ |> map HOLogic.mk_Trueprop
+ end
+ | NONE => []
+ end
+(* Proof.context -> string * typ list -> term list *)
+fun optimized_quot_type_axioms ctxt stds abs_z =
+ let
+ val thy = ProofContext.theory_of ctxt
+ val abs_T = Type abs_z
+ val rep_T = rep_type_for_quot_type thy abs_T
+ val equiv_rel = equiv_relation_for_quot_type thy abs_T
+ val a_var = Var (("a", 0), abs_T)
+ val x_var = Var (("x", 0), rep_T)
+ val y_var = Var (("y", 0), rep_T)
+ val x = (@{const_name Quot}, rep_T --> abs_T)
+ val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
+ val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
+ val normal_x = normal_t $ x_var
+ val normal_y = normal_t $ y_var
+ val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
+ in
+ [Logic.mk_equals (normal_t $ sel_a_t, sel_a_t),
+ Logic.list_implies
+ ([@{const Not} $ (is_unknown_t $ normal_x),
+ @{const Not} $ (is_unknown_t $ normal_y),
+ equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
+ Logic.mk_equals (normal_x, normal_y)),
+ Logic.list_implies
+ ([HOLogic.mk_Trueprop (@{const Not} $ (is_unknown_t $ normal_x)),
+ HOLogic.mk_Trueprop (@{const Not} $ HOLogic.mk_eq (normal_x, x_var))],
+ HOLogic.mk_Trueprop (equiv_rel $ x_var $ normal_x))]
+ end
+
(* hol_context -> typ -> term list *)
fun codatatype_bisim_axioms (hol_ctxt as {thy, stds, ...}) T =
let
@@ -2155,7 +2151,7 @@
end
fun inductive_pred_axiom hol_ctxt (x as (s, T)) =
if String.isPrefix ubfp_prefix s orelse String.isPrefix lbfp_prefix s then
- let val x' = (after_name_sep s, T) in
+ let val x' = (strip_first_name_sep s |> snd, T) in
raw_inductive_pred_axiom hol_ctxt x' |> subst_atomic [(Const x', Const x)]
end
else
@@ -2177,6 +2173,8 @@
strip_comb t1 |> snd |> forall is_Var
| _ => false
+(** Type preprocessing **)
+
(* term list -> term list *)
fun merge_type_vars_in_terms ts =
let
@@ -2222,198 +2220,4 @@
fun ground_types_in_terms hol_ctxt binarize ts =
fold (fold_types (add_ground_types hol_ctxt binarize)) ts []
-(* theory -> const_table -> styp -> int list *)
-fun const_format thy def_table (x as (s, T)) =
- if String.isPrefix unrolled_prefix s then
- const_format thy def_table (original_name s, range_type T)
- else if String.isPrefix skolem_prefix s then
- let
- val k = unprefix skolem_prefix s
- |> strip_first_name_sep |> fst |> space_explode "@"
- |> hd |> Int.fromString |> the
- in [k, num_binder_types T - k] end
- else if original_name s <> s then
- [num_binder_types T]
- else case def_of_const thy def_table x of
- SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then
- let val k = length (strip_abs_vars t') in
- [k, num_binder_types T - k]
- end
- else
- [num_binder_types T]
- | NONE => [num_binder_types T]
-(* int list -> int list -> int list *)
-fun intersect_formats _ [] = []
- | intersect_formats [] _ = []
- | intersect_formats ks1 ks2 =
- let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in
- intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else []))
- (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @
- [Int.min (k1, k2)]
- end
-
-(* theory -> const_table -> (term option * int list) list -> term -> int list *)
-fun lookup_format thy def_table formats t =
- case AList.lookup (fn (SOME x, SOME y) =>
- (term_match thy) (x, y) | _ => false)
- formats (SOME t) of
- SOME format => format
- | NONE => let val format = the (AList.lookup (op =) formats NONE) in
- case t of
- Const x => intersect_formats format
- (const_format thy def_table x)
- | _ => format
- end
-
-(* int list -> int list -> typ -> typ *)
-fun format_type default_format format T =
- let
- val T = uniterize_unarize_unbox_etc_type T
- val format = format |> filter (curry (op <) 0)
- in
- if forall (curry (op =) 1) format then
- T
- else
- let
- val (binder_Ts, body_T) = strip_type T
- val batched =
- binder_Ts
- |> map (format_type default_format default_format)
- |> rev |> chunk_list_unevenly (rev format)
- |> map (HOLogic.mk_tupleT o rev)
- in List.foldl (op -->) body_T batched end
- end
-(* theory -> const_table -> (term option * int list) list -> term -> typ *)
-fun format_term_type thy def_table formats t =
- format_type (the (AList.lookup (op =) formats NONE))
- (lookup_format thy def_table formats t) (fastype_of t)
-
-(* int list -> int -> int list -> int list *)
-fun repair_special_format js m format =
- m - 1 downto 0 |> chunk_list_unevenly (rev format)
- |> map (rev o filter_out (member (op =) js))
- |> filter_out null |> map length |> rev
-
-(* hol_context -> string * string -> (term option * int list) list
- -> styp -> term * typ *)
-fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...}
- : hol_context) (base_name, step_name) formats =
- let
- val default_format = the (AList.lookup (op =) formats NONE)
- (* styp -> term * typ *)
- fun do_const (x as (s, T)) =
- (if String.isPrefix special_prefix s then
- let
- (* term -> term *)
- val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t')
- val (x' as (_, T'), js, ts) =
- AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T)
- |> the_single
- val max_j = List.last js
- val Ts = List.take (binder_types T', max_j + 1)
- val missing_js = filter_out (member (op =) js) (0 upto max_j)
- val missing_Ts = filter_indices missing_js Ts
- (* int -> indexname *)
- fun nth_missing_var n =
- ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
- val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
- val vars = special_bounds ts @ missing_vars
- val ts' =
- map (fn j =>
- case AList.lookup (op =) (js ~~ ts) j of
- SOME t => do_term t
- | NONE =>
- Var (nth missing_vars
- (find_index (curry (op =) j) missing_js)))
- (0 upto max_j)
- val t = do_const x' |> fst
- val format =
- case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
- | _ => false) formats (SOME t) of
- SOME format =>
- repair_special_format js (num_binder_types T') format
- | NONE =>
- const_format thy def_table x'
- |> repair_special_format js (num_binder_types T')
- |> intersect_formats default_format
- in
- (list_comb (t, ts') |> fold_rev abs_var vars,
- format_type default_format format T)
- end
- else if String.isPrefix uncurry_prefix s then
- let
- val (ss, s') = unprefix uncurry_prefix s
- |> strip_first_name_sep |>> space_explode "@"
- in
- if String.isPrefix step_prefix s' then
- do_const (s', T)
- else
- let
- val k = the (Int.fromString (hd ss))
- val j = the (Int.fromString (List.last ss))
- val (before_Ts, (tuple_T, rest_T)) =
- strip_n_binders j T ||> (strip_n_binders 1 #>> hd)
- val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T
- in do_const (s', T') end
- end
- else if String.isPrefix unrolled_prefix s then
- let val t = Const (original_name s, range_type T) in
- (lambda (Free (iter_var_prefix, nat_T)) t,
- format_type default_format
- (lookup_format thy def_table formats t) T)
- end
- else if String.isPrefix base_prefix s then
- (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T),
- format_type default_format default_format T)
- else if String.isPrefix step_prefix s then
- (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
- format_type default_format default_format T)
- else if String.isPrefix quot_normal_prefix s then
- let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
- (t, format_term_type thy def_table formats t)
- end
- else if String.isPrefix skolem_prefix s then
- let
- val ss = the (AList.lookup (op =) (!skolems) s)
- val (Ts, Ts') = chop (length ss) (binder_types T)
- val frees = map Free (ss ~~ Ts)
- val s' = original_name s
- in
- (fold lambda frees (Const (s', Ts' ---> T)),
- format_type default_format
- (lookup_format thy def_table formats (Const x)) T)
- end
- else if String.isPrefix eval_prefix s then
- let
- val t = nth evals (the (Int.fromString (unprefix eval_prefix s)))
- in (t, format_term_type thy def_table formats t) end
- else if s = @{const_name undefined_fast_The} then
- (Const (nitpick_prefix ^ "The fallback", T),
- format_type default_format
- (lookup_format thy def_table formats
- (Const (@{const_name The}, (T --> bool_T) --> T))) T)
- else if s = @{const_name undefined_fast_Eps} then
- (Const (nitpick_prefix ^ "Eps fallback", T),
- format_type default_format
- (lookup_format thy def_table formats
- (Const (@{const_name Eps}, (T --> bool_T) --> T))) T)
- else
- let val t = Const (original_name s, T) in
- (t, format_term_type thy def_table formats t)
- end)
- |>> map_types uniterize_unarize_unbox_etc_type
- |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
- in do_const end
-
-(* styp -> string *)
-fun assign_operator_for_const (s, T) =
- if String.isPrefix ubfp_prefix s then
- if is_fun_type T then "\<subseteq>" else "\<le>"
- else if String.isPrefix lbfp_prefix s then
- if is_fun_type T then "\<supseteq>" else "\<ge>"
- else if original_name s <> s then
- assign_operator_for_const (after_name_sep s, T)
- else
- "="
-
end;
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Thu Mar 11 19:06:03 2010 +0100
@@ -16,15 +16,20 @@
show_skolems: bool,
show_datatypes: bool,
show_consts: bool}
- type term_postprocessor = Proof.context -> typ -> term -> term
+ type term_postprocessor =
+ Proof.context -> string -> (typ -> term list) -> typ -> term -> term
structure NameTable : TABLE
+ val irrelevant : string
+ val unknown : string
+ val unrep : string
val register_term_postprocessor :
typ -> term_postprocessor -> theory -> theory
val unregister_term_postprocessor : typ -> theory -> theory
val tuple_list_for_name :
nut NameTable.table -> Kodkod.raw_bound list -> nut -> int list list
+ val dest_plain_fun : term -> bool * (term list * term list)
val reconstruct_hol_model :
params -> scope -> (term option * int list) list -> styp list -> nut list
-> nut list -> nut list -> nut NameTable.table -> Kodkod.raw_bound list
@@ -51,7 +56,8 @@
show_datatypes: bool,
show_consts: bool}
-type term_postprocessor = Proof.context -> typ -> term -> term
+type term_postprocessor =
+ Proof.context -> string -> (typ -> term list) -> typ -> term -> term
structure Data = Theory_Data(
type T = (typ * term_postprocessor) list
@@ -59,12 +65,14 @@
val extend = I
fun merge (ps1, ps2) = AList.merge (op =) (K true) (ps1, ps2))
+val irrelevant = "_"
val unknown = "?"
val unrep = "\<dots>"
val maybe_mixfix = "_\<^sup>?"
val base_mixfix = "_\<^bsub>base\<^esub>"
val step_mixfix = "_\<^bsub>step\<^esub>"
val abs_mixfix = "\<guillemotleft>_\<guillemotright>"
+val arg_var_prefix = "x"
val cyclic_co_val_name = "\<omega>"
val cyclic_const_prefix = "\<xi>"
val cyclic_type_name = nitpick_prefix ^ cyclic_const_prefix
@@ -73,6 +81,31 @@
type atom_pool = ((string * int) * int list) list
+(* Proof.context -> ((string * string) * (string * string)) * Proof.context *)
+fun add_wacky_syntax ctxt =
+ let
+ (* term -> string *)
+ val name_of = fst o dest_Const
+ val thy = ProofContext.theory_of ctxt |> Context.reject_draft
+ val (maybe_t, thy) =
+ Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
+ Mixfix (maybe_mixfix, [1000], 1000)) thy
+ val (abs_t, thy) =
+ Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
+ Mixfix (abs_mixfix, [40], 40)) thy
+ val (base_t, thy) =
+ Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
+ Mixfix (base_mixfix, [1000], 1000)) thy
+ val (step_t, thy) =
+ Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
+ Mixfix (step_mixfix, [1000], 1000)) thy
+ in
+ (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
+ ProofContext.transfer_syntax thy ctxt)
+ end
+
+(** Term reconstruction **)
+
(* atom_pool Unsynchronized.ref -> string -> int -> int -> string *)
fun nth_atom_suffix pool s j k =
(case AList.lookup (op =) (!pool) (s, k) of
@@ -182,9 +215,9 @@
(* typ -> typ -> (term * term) list -> term *)
fun aux T1 T2 [] =
Const (if maybe_opt then opt_flag else non_opt_flag, T1 --> T2)
- | aux T1 T2 ((t1, t2) :: ps) =
+ | aux T1 T2 ((t1, t2) :: tps) =
Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
- $ aux T1 T2 ps $ t1 $ t2
+ $ aux T1 T2 tps $ t1 $ t2
in aux T1 T2 o rev end
(* term -> bool *)
fun is_plain_fun (Const (s, _)) = (s = opt_flag orelse s = non_opt_flag)
@@ -195,9 +228,12 @@
val dest_plain_fun =
let
(* term -> bool * (term list * term list) *)
- fun aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
+ fun aux (Abs (_, _, Const (s, _))) = (s <> irrelevant, ([], []))
+ | aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
| aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
- let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end
+ let val (maybe_opt, (ts1, ts2)) = aux t0 in
+ (maybe_opt, (t1 :: ts1, t2 :: ts2))
+ end
| aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
in apsnd (pairself rev) o aux end
@@ -227,22 +263,22 @@
(* typ -> typ -> typ -> typ -> term -> term *)
fun do_curry T1 T1a T1b T2 t =
let
- val (maybe_opt, ps) = dest_plain_fun t
- val ps =
- ps |>> map (break_in_two T1 T1a T1b)
- |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
- |> AList.coalesce (op =)
- |> map (apsnd (make_plain_fun maybe_opt T1b T2))
- in make_plain_fun maybe_opt T1a (T1b --> T2) ps end
+ val (maybe_opt, tsp) = dest_plain_fun t
+ val tps =
+ tsp |>> map (break_in_two T1 T1a T1b)
+ |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
+ |> AList.coalesce (op =)
+ |> map (apsnd (make_plain_fun maybe_opt T1b T2))
+ in make_plain_fun maybe_opt T1a (T1b --> T2) tps end
(* typ -> typ -> term -> term *)
and do_uncurry T1 T2 t =
let
val (maybe_opt, tsp) = dest_plain_fun t
- val ps =
+ val tps =
tsp |> op ~~
|> maps (fn (t1, t2) =>
multi_pair_up T1 t1 (snd (dest_plain_fun t2)))
- in make_plain_fun maybe_opt T1 T2 ps end
+ in make_plain_fun maybe_opt T1 T2 tps end
(* typ -> typ -> typ -> typ -> term -> term *)
and do_arrow T1' T2' _ _ (Const (s, _)) = Const (s, T1' --> T2')
| do_arrow T1' T2' T1 T2
@@ -291,22 +327,37 @@
(* theory -> typ * typ -> bool *)
fun varified_type_match thy (candid_T, pat_T) =
strict_type_match thy (candid_T, Logic.varifyT pat_T)
-(* Proof.context -> typ -> term -> term *)
-fun postprocess_term ctxt T =
- let val thy = ProofContext.theory_of ctxt in
- if null (Data.get thy) then
- I
- else
- (T |> AList.lookup (varified_type_match thy) (Data.get thy)
- |> the_default (K (K I))) ctxt T
- end
+(* atom_pool -> (string * string) * (string * string) -> scope -> nut list
+ -> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list -> typ
+ -> term list *)
+fun all_values_of_type pool wacky_names (scope as {card_assigns, ...} : scope)
+ sel_names rel_table bounds card T =
+ let
+ val card = if card = 0 then card_of_type card_assigns T else card
+ (* nat -> term *)
+ fun nth_value_of_type n =
+ let
+ (* bool -> term *)
+ fun term unfold =
+ reconstruct_term unfold pool wacky_names scope sel_names rel_table
+ bounds T T (Atom (card, 0)) [[n]]
+ in
+ case term false of
+ t as Const (s, _) =>
+ if String.isPrefix cyclic_const_prefix s then
+ HOLogic.mk_eq (t, term true)
+ else
+ t
+ | t => t
+ end
+ in index_seq 0 card |> map nth_value_of_type |> sort nice_term_ord end
(* bool -> atom_pool -> (string * string) * (string * string) -> scope
-> nut list -> nut list -> nut list -> nut NameTable.table
-> KK.raw_bound list -> typ -> typ -> rep -> int list list -> term *)
-fun reconstruct_term unfold pool ((maybe_name, abs_name), _)
- ({hol_ctxt as {ctxt, thy, stds, ...}, binarize, card_assigns, bits,
- datatypes, ofs, ...} : scope) sel_names rel_table bounds =
+and reconstruct_term unfold pool (wacky_names as ((maybe_name, abs_name), _))
+ (scope as {hol_ctxt as {ctxt, thy, stds, ...}, binarize, card_assigns,
+ bits, datatypes, ofs, ...}) sel_names rel_table bounds =
let
val for_auto = (maybe_name = "")
(* int list list -> int *)
@@ -318,6 +369,25 @@
fold (fn j => Integer.add (reasonable_power 2 j |> j = bits ? op ~))
js 0
end
+ (* typ -> term list *)
+ val all_values =
+ all_values_of_type pool wacky_names scope sel_names rel_table bounds 0
+ (* typ -> term -> term *)
+ fun postprocess_term (Type (@{type_name fun}, _)) = I
+ | postprocess_term T =
+ if null (Data.get thy) then
+ I
+ else case AList.lookup (varified_type_match thy) (Data.get thy) T of
+ SOME postproc => postproc ctxt maybe_name all_values T
+ | NONE => I
+ (* typ list -> term -> term *)
+ fun postprocess_subterms Ts (t1 $ t2) =
+ let val t = postprocess_subterms Ts t1 $ postprocess_subterms Ts t2 in
+ postprocess_term (fastype_of1 (Ts, t)) t
+ end
+ | postprocess_subterms Ts (Abs (s, T, t')) =
+ Abs (s, T, postprocess_subterms (T :: Ts) t')
+ | postprocess_subterms Ts t = postprocess_term (fastype_of1 (Ts, t)) t
(* bool -> typ -> typ -> (term * term) list -> term *)
fun make_set maybe_opt T1 T2 tps =
let
@@ -352,16 +422,16 @@
(T1 --> T2) --> T1 --> T2 --> T1 --> T2)
(* (term * term) list -> term *)
fun aux' [] = Const (@{const_abbrev Map.empty}, T1 --> T2)
- | aux' ((t1, t2) :: ps) =
+ | aux' ((t1, t2) :: tps) =
(case t2 of
- Const (@{const_name None}, _) => aux' ps
- | _ => update_const $ aux' ps $ t1 $ t2)
- fun aux ps =
+ Const (@{const_name None}, _) => aux' tps
+ | _ => update_const $ aux' tps $ t1 $ t2)
+ fun aux tps =
if maybe_opt andalso not (is_complete_type datatypes false T1) then
- update_const $ aux' ps $ Const (unrep, T1)
+ update_const $ aux' tps $ Const (unrep, T1)
$ (Const (@{const_name Some}, T2' --> T2) $ Const (unknown, T2'))
else
- aux' ps
+ aux' tps
in aux end
(* typ list -> term -> term *)
fun polish_funs Ts t =
@@ -396,7 +466,11 @@
| Abs (s, T, t') => Abs (s, T, polish_funs (T :: Ts) t')
| Const (s, Type (@{type_name fun}, [T1, T2])) =>
if s = opt_flag orelse s = non_opt_flag then
- Abs ("x", T1, Const (unknown, T2))
+ Abs ("x", T1,
+ Const (if is_complete_type datatypes false T1 then
+ irrelevant
+ else
+ unknown, T2))
else
t
| t => t
@@ -541,7 +615,6 @@
t
end
end
- |> postprocess_term ctxt T
(* (typ * int) list -> int -> rep -> typ -> typ -> typ -> int list
-> term *)
and term_for_vect seen k R T1 T2 T' js =
@@ -594,7 +667,215 @@
raise ARG ("Nitpick_Model.reconstruct_term.term_for_rep",
Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
string_of_int (length jss))
- in polish_funs [] o unarize_unbox_etc_term oooo term_for_rep true [] end
+ in
+ postprocess_subterms [] o polish_funs [] o unarize_unbox_etc_term
+ oooo term_for_rep true []
+ end
+
+(** Constant postprocessing **)
+
+(* int -> typ -> typ list *)
+fun dest_n_tuple_type 1 T = [T]
+ | dest_n_tuple_type n (Type (_, [T1, T2])) =
+ T1 :: dest_n_tuple_type (n - 1) T2
+ | dest_n_tuple_type _ T =
+ raise TYPE ("Nitpick_Model.dest_n_tuple_type", [T], [])
+
+(* theory -> const_table -> styp -> int list *)
+fun const_format thy def_table (x as (s, T)) =
+ if String.isPrefix unrolled_prefix s then
+ const_format thy def_table (original_name s, range_type T)
+ else if String.isPrefix skolem_prefix s then
+ let
+ val k = unprefix skolem_prefix s
+ |> strip_first_name_sep |> fst |> space_explode "@"
+ |> hd |> Int.fromString |> the
+ in [k, num_binder_types T - k] end
+ else if original_name s <> s then
+ [num_binder_types T]
+ else case def_of_const thy def_table x of
+ SOME t' => if fixpoint_kind_of_rhs t' <> NoFp then
+ let val k = length (strip_abs_vars t') in
+ [k, num_binder_types T - k]
+ end
+ else
+ [num_binder_types T]
+ | NONE => [num_binder_types T]
+(* int list -> int list -> int list *)
+fun intersect_formats _ [] = []
+ | intersect_formats [] _ = []
+ | intersect_formats ks1 ks2 =
+ let val ((ks1', k1), (ks2', k2)) = pairself split_last (ks1, ks2) in
+ intersect_formats (ks1' @ (if k1 > k2 then [k1 - k2] else []))
+ (ks2' @ (if k2 > k1 then [k2 - k1] else [])) @
+ [Int.min (k1, k2)]
+ end
+
+(* theory -> const_table -> (term option * int list) list -> term -> int list *)
+fun lookup_format thy def_table formats t =
+ case AList.lookup (fn (SOME x, SOME y) =>
+ (term_match thy) (x, y) | _ => false)
+ formats (SOME t) of
+ SOME format => format
+ | NONE => let val format = the (AList.lookup (op =) formats NONE) in
+ case t of
+ Const x => intersect_formats format
+ (const_format thy def_table x)
+ | _ => format
+ end
+
+(* int list -> int list -> typ -> typ *)
+fun format_type default_format format T =
+ let
+ val T = uniterize_unarize_unbox_etc_type T
+ val format = format |> filter (curry (op <) 0)
+ in
+ if forall (curry (op =) 1) format then
+ T
+ else
+ let
+ val (binder_Ts, body_T) = strip_type T
+ val batched =
+ binder_Ts
+ |> map (format_type default_format default_format)
+ |> rev |> chunk_list_unevenly (rev format)
+ |> map (HOLogic.mk_tupleT o rev)
+ in List.foldl (op -->) body_T batched end
+ end
+(* theory -> const_table -> (term option * int list) list -> term -> typ *)
+fun format_term_type thy def_table formats t =
+ format_type (the (AList.lookup (op =) formats NONE))
+ (lookup_format thy def_table formats t) (fastype_of t)
+
+(* int list -> int -> int list -> int list *)
+fun repair_special_format js m format =
+ m - 1 downto 0 |> chunk_list_unevenly (rev format)
+ |> map (rev o filter_out (member (op =) js))
+ |> filter_out null |> map length |> rev
+
+(* hol_context -> string * string -> (term option * int list) list
+ -> styp -> term * typ *)
+fun user_friendly_const ({thy, evals, def_table, skolems, special_funs, ...}
+ : hol_context) (base_name, step_name) formats =
+ let
+ val default_format = the (AList.lookup (op =) formats NONE)
+ (* styp -> term * typ *)
+ fun do_const (x as (s, T)) =
+ (if String.isPrefix special_prefix s then
+ let
+ (* term -> term *)
+ val do_term = map_aterms (fn Const x => fst (do_const x) | t' => t')
+ val (x' as (_, T'), js, ts) =
+ AList.find (op =) (!special_funs) (s, unarize_unbox_etc_type T)
+ |> the_single
+ val max_j = List.last js
+ val Ts = List.take (binder_types T', max_j + 1)
+ val missing_js = filter_out (member (op =) js) (0 upto max_j)
+ val missing_Ts = filter_indices missing_js Ts
+ (* int -> indexname *)
+ fun nth_missing_var n =
+ ((arg_var_prefix ^ nat_subscript (n + 1), 0), nth missing_Ts n)
+ val missing_vars = map nth_missing_var (0 upto length missing_js - 1)
+ val vars = special_bounds ts @ missing_vars
+ val ts' =
+ map (fn j =>
+ case AList.lookup (op =) (js ~~ ts) j of
+ SOME t => do_term t
+ | NONE =>
+ Var (nth missing_vars
+ (find_index (curry (op =) j) missing_js)))
+ (0 upto max_j)
+ val t = do_const x' |> fst
+ val format =
+ case AList.lookup (fn (SOME t1, SOME t2) => term_match thy (t1, t2)
+ | _ => false) formats (SOME t) of
+ SOME format =>
+ repair_special_format js (num_binder_types T') format
+ | NONE =>
+ const_format thy def_table x'
+ |> repair_special_format js (num_binder_types T')
+ |> intersect_formats default_format
+ in
+ (list_comb (t, ts') |> fold_rev abs_var vars,
+ format_type default_format format T)
+ end
+ else if String.isPrefix uncurry_prefix s then
+ let
+ val (ss, s') = unprefix uncurry_prefix s
+ |> strip_first_name_sep |>> space_explode "@"
+ in
+ if String.isPrefix step_prefix s' then
+ do_const (s', T)
+ else
+ let
+ val k = the (Int.fromString (hd ss))
+ val j = the (Int.fromString (List.last ss))
+ val (before_Ts, (tuple_T, rest_T)) =
+ strip_n_binders j T ||> (strip_n_binders 1 #>> hd)
+ val T' = before_Ts ---> dest_n_tuple_type k tuple_T ---> rest_T
+ in do_const (s', T') end
+ end
+ else if String.isPrefix unrolled_prefix s then
+ let val t = Const (original_name s, range_type T) in
+ (lambda (Free (iter_var_prefix, nat_T)) t,
+ format_type default_format
+ (lookup_format thy def_table formats t) T)
+ end
+ else if String.isPrefix base_prefix s then
+ (Const (base_name, T --> T) $ Const (unprefix base_prefix s, T),
+ format_type default_format default_format T)
+ else if String.isPrefix step_prefix s then
+ (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
+ format_type default_format default_format T)
+ else if String.isPrefix quot_normal_prefix s then
+ let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
+ (t, format_term_type thy def_table formats t)
+ end
+ else if String.isPrefix skolem_prefix s then
+ let
+ val ss = the (AList.lookup (op =) (!skolems) s)
+ val (Ts, Ts') = chop (length ss) (binder_types T)
+ val frees = map Free (ss ~~ Ts)
+ val s' = original_name s
+ in
+ (fold lambda frees (Const (s', Ts' ---> T)),
+ format_type default_format
+ (lookup_format thy def_table formats (Const x)) T)
+ end
+ else if String.isPrefix eval_prefix s then
+ let
+ val t = nth evals (the (Int.fromString (unprefix eval_prefix s)))
+ in (t, format_term_type thy def_table formats t) end
+ else if s = @{const_name undefined_fast_The} then
+ (Const (nitpick_prefix ^ "The fallback", T),
+ format_type default_format
+ (lookup_format thy def_table formats
+ (Const (@{const_name The}, (T --> bool_T) --> T))) T)
+ else if s = @{const_name undefined_fast_Eps} then
+ (Const (nitpick_prefix ^ "Eps fallback", T),
+ format_type default_format
+ (lookup_format thy def_table formats
+ (Const (@{const_name Eps}, (T --> bool_T) --> T))) T)
+ else
+ let val t = Const (original_name s, T) in
+ (t, format_term_type thy def_table formats t)
+ end)
+ |>> map_types uniterize_unarize_unbox_etc_type
+ |>> shorten_names_in_term |>> Term.map_abs_vars shortest_name
+ in do_const end
+
+(* styp -> string *)
+fun assign_operator_for_const (s, T) =
+ if String.isPrefix ubfp_prefix s then
+ if is_fun_type T then "\<subseteq>" else "\<le>"
+ else if String.isPrefix lbfp_prefix s then
+ if is_fun_type T then "\<supseteq>" else "\<ge>"
+ else if original_name s <> s then
+ assign_operator_for_const (strip_first_name_sep s |> snd, T)
+ else
+ "="
+
+(** Model reconstruction **)
(* atom_pool -> scope -> nut list -> nut NameTable.table -> KK.raw_bound list
-> nut -> term *)
@@ -605,29 +886,6 @@
rel_table bounds T T (rep_of name)
end
-(* Proof.context -> ((string * string) * (string * string)) * Proof.context *)
-fun add_wacky_syntax ctxt =
- let
- (* term -> string *)
- val name_of = fst o dest_Const
- val thy = ProofContext.theory_of ctxt |> Context.reject_draft
- val (maybe_t, thy) =
- Sign.declare_const ((@{binding nitpick_maybe}, @{typ "'a => 'a"}),
- Mixfix (maybe_mixfix, [1000], 1000)) thy
- val (abs_t, thy) =
- Sign.declare_const ((@{binding nitpick_abs}, @{typ "'a => 'b"}),
- Mixfix (abs_mixfix, [40], 40)) thy
- val (base_t, thy) =
- Sign.declare_const ((@{binding nitpick_base}, @{typ "'a => 'a"}),
- Mixfix (base_mixfix, [1000], 1000)) thy
- val (step_t, thy) =
- Sign.declare_const ((@{binding nitpick_step}, @{typ "'a => 'a"}),
- Mixfix (step_mixfix, [1000], 1000)) thy
- in
- (pairself (pairself name_of) ((maybe_t, abs_t), (base_t, step_t)),
- ProofContext.transfer_syntax thy ctxt)
- end
-
(* term -> term *)
fun unfold_outer_the_binders (t as Const (@{const_name The}, _)
$ Abs (s, T, Const (@{const_name "op ="}, _)
@@ -702,14 +960,14 @@
t
| t => t
end
- (* nat -> typ -> typ list *)
- fun all_values_of_type card T =
- index_seq 0 card |> map (nth_value_of_type card T) |> sort nice_term_ord
+ (* nat -> typ -> term list *)
+ val all_values =
+ all_values_of_type pool wacky_names scope sel_names rel_table bounds
(* dtype_spec list -> dtype_spec -> bool *)
fun is_codatatype_wellformed (cos : dtype_spec list)
({typ, card, ...} : dtype_spec) =
let
- val ts = all_values_of_type card typ
+ val ts = all_values card typ
val max_depth = Integer.sum (map #card cos)
in
forall (not o bisimilar_values (map #typ cos) max_depth)
@@ -750,7 +1008,7 @@
| _ => []) @
[Pretty.str "=",
Pretty.enum "," "{" "}"
- (map (Syntax.pretty_term ctxt) (all_values_of_type card typ) @
+ (map (Syntax.pretty_term ctxt) (all_values card typ) @
(if fun_from_pair complete false then []
else [Pretty.str unrep]))]))
(* typ -> dtype_spec list *)
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML Thu Mar 11 19:06:03 2010 +0100
@@ -33,27 +33,40 @@
binary coding. *)
val binary_int_threshold = 3
+(* bool -> term -> bool *)
+val may_use_binary_ints =
+ let
+ (* bool -> term -> bool *)
+ fun aux def (Const (@{const_name "=="}, _) $ t1 $ t2) =
+ aux def t1 andalso aux false t2
+ | aux def (@{const "==>"} $ t1 $ t2) = aux false t1 andalso aux def t2
+ | aux def (Const (@{const_name "op ="}, _) $ t1 $ t2) =
+ aux def t1 andalso aux false t2
+ | aux def (@{const "op -->"} $ t1 $ t2) = aux false t1 andalso aux def t2
+ | aux def (t1 $ t2) = aux def t1 andalso aux def t2
+ | aux def (t as Const (s, _)) =
+ (not def orelse t <> @{const Suc}) andalso
+ not (member (op =) [@{const_name Abs_Frac}, @{const_name Rep_Frac},
+ @{const_name nat_gcd}, @{const_name nat_lcm},
+ @{const_name Frac}, @{const_name norm_frac}] s)
+ | aux def (Abs (_, _, t')) = aux def t'
+ | aux _ _ = true
+ in aux end
(* term -> bool *)
-fun may_use_binary_ints (t1 $ t2) =
- may_use_binary_ints t1 andalso may_use_binary_ints t2
- | may_use_binary_ints (t as Const (s, _)) =
- t <> @{const Suc} andalso
- not (member (op =) [@{const_name Abs_Frac}, @{const_name Rep_Frac},
- @{const_name nat_gcd}, @{const_name nat_lcm},
- @{const_name Frac}, @{const_name norm_frac}] s)
- | may_use_binary_ints (Abs (_, _, t')) = may_use_binary_ints t'
- | may_use_binary_ints _ = true
-fun should_use_binary_ints (t1 $ t2) =
- should_use_binary_ints t1 orelse should_use_binary_ints t2
- | should_use_binary_ints (Const (s, T)) =
- ((s = @{const_name times} orelse s = @{const_name div}) andalso
- is_integer_type (body_type T)) orelse
- (String.isPrefix numeral_prefix s andalso
- let val n = the (Int.fromString (unprefix numeral_prefix s)) in
- n < ~ binary_int_threshold orelse n > binary_int_threshold
- end)
- | should_use_binary_ints (Abs (_, _, t')) = should_use_binary_ints t'
- | should_use_binary_ints _ = false
+val should_use_binary_ints =
+ let
+ (* term -> bool *)
+ fun aux (t1 $ t2) = aux t1 orelse aux t2
+ | aux (Const (s, T)) =
+ ((s = @{const_name times} orelse s = @{const_name div}) andalso
+ is_integer_type (body_type T)) orelse
+ (String.isPrefix numeral_prefix s andalso
+ let val n = the (Int.fromString (unprefix numeral_prefix s)) in
+ n < ~ binary_int_threshold orelse n > binary_int_threshold
+ end)
+ | aux (Abs (_, _, t')) = aux t'
+ | aux _ = false
+ in aux end
(** Uncurrying **)
@@ -1369,7 +1382,8 @@
is_standard_datatype thy stds nat_T andalso
case binary_ints of
SOME false => false
- | _ => forall may_use_binary_ints (nondef_ts @ def_ts) andalso
+ | _ => forall (may_use_binary_ints false) nondef_ts andalso
+ forall (may_use_binary_ints true) def_ts andalso
(binary_ints = SOME true orelse
exists should_use_binary_ints (nondef_ts @ def_ts))
val box = exists (not_equal (SOME false) o snd) boxes
--- a/src/HOL/Wellfounded.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/Wellfounded.thy Thu Mar 11 19:06:03 2010 +0100
@@ -8,7 +8,7 @@
header {*Well-founded Recursion*}
theory Wellfounded
-imports Finite_Set Transitive_Closure
+imports Transitive_Closure
uses ("Tools/Function/size.ML")
begin
--- a/src/HOL/ex/Summation.thy Thu Mar 11 19:05:46 2010 +0100
+++ b/src/HOL/ex/Summation.thy Thu Mar 11 19:06:03 2010 +0100
@@ -10,7 +10,7 @@
lemma add_setsum_orient:
"setsum f {k..<j} + setsum f {l..<k} = setsum f {l..<k} + setsum f {k..<j}"
- by (fact plus.commute)
+ by (fact add.commute)
lemma add_setsum_int:
fixes j k l :: int
--- a/src/Pure/General/sha1_polyml.ML Thu Mar 11 19:05:46 2010 +0100
+++ b/src/Pure/General/sha1_polyml.ML Thu Mar 11 19:06:03 2010 +0100
@@ -15,15 +15,21 @@
let val c = CInterface.fromCchar (CInterface.offset i CInterface.Cchar arr)
in (op ^) (pairself hex_digit (Integer.div_mod (Char.ord c) 16)) end
+val lib_path =
+ ("$ML_HOME/" ^ (if String.isSuffix "cygwin" ml_system then "sha1.dll" else "libsha1.so"))
+ |> Path.explode;
+
fun digest_external str =
let
val digest = CInterface.alloc 20 CInterface.Cchar;
- val _ = CInterface.call3 (CInterface.get_sym "sha1" "sha1_buffer")
- (CInterface.STRING, CInterface.INT, CInterface.POINTER)
- CInterface.POINTER (str, size str, CInterface.address digest);
+ val _ =
+ CInterface.call3 (CInterface.get_sym (File.platform_path lib_path) "sha1_buffer")
+ (CInterface.STRING, CInterface.INT, CInterface.POINTER)
+ CInterface.POINTER (str, size str, CInterface.address digest);
in fold (suffix o hex_string digest) (0 upto 19) "" end;
fun digest str = digest_external str
- handle CInterface.Foreign _ => SHA1.digest str;
+ handle CInterface.Foreign msg =>
+ (warning (msg ^ "\nUsing slow ML implementation of SHA1.digest"); SHA1.digest str);
end;