author | blanchet |
Fri, 03 Sep 2010 13:54:39 +0200 | |
changeset 39113 | 91ba394cc525 |
parent 39105 | 3b9e020c3908 (diff) |
parent 39112 | 611e41ef07c3 (current diff) |
child 39114 | 240e2b41a041 |
--- a/NEWS Fri Sep 03 13:54:04 2010 +0200 +++ b/NEWS Fri Sep 03 13:54:39 2010 +0200 @@ -37,8 +37,9 @@ Thy_Output.break thy_output_break show_question_marks show_question_marks - -Note that the corresponding "..._default" references may be only + show_consts show_consts + +Note that the corresponding "..._default" references in ML may be only changed globally at the ROOT session setup, but *not* within a theory. @@ -200,6 +201,9 @@ derive instantiated and simplified equations for inductive predicates, similar to inductive_cases. +* "bij f" is now an abbreviation of "bij_betw f UNIV UNIV". surj_on is a +generalization of surj, and "surj f" is now a abbreviation of "surj_on f UNIV". +The theorems bij_def and surj_def are unchanged. *** FOL ***
--- a/doc-src/Classes/Thy/document/Classes.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Classes/Thy/document/Classes.tex Fri Sep 03 13:54:39 2010 +0200 @@ -1134,26 +1134,27 @@ \noindent% \hspace*{0pt}module Example where {\char123}\\ \hspace*{0pt}\\ -\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Example.Nat;\\ \hspace*{0pt}\\ -\hspace*{0pt}nat{\char95}aux ::~Integer -> Nat -> Nat;\\ -\hspace*{0pt}nat{\char95}aux i n = (if i <= 0 then n else nat{\char95}aux (i - 1) (Suc n));\\ +\hspace*{0pt}nat{\char95}aux ::~Integer -> Example.Nat -> Example.Nat;\\ +\hspace*{0pt}nat{\char95}aux i n =\\ +\hspace*{0pt} ~(if i <= 0 then n else Example.nat{\char95}aux (i - 1) (Example.Suc n));\\ \hspace*{0pt}\\ -\hspace*{0pt}nat ::~Integer -> Nat;\\ -\hspace*{0pt}nat i = nat{\char95}aux i Zero{\char95}nat;\\ +\hspace*{0pt}nat ::~Integer -> Example.Nat;\\ +\hspace*{0pt}nat i = Example.nat{\char95}aux i Example.Zero{\char95}nat;\\ \hspace*{0pt}\\ \hspace*{0pt}class Semigroup a where {\char123}\\ \hspace*{0pt} ~mult ::~a -> a -> a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoidl a where {\char123}\\ +\hspace*{0pt}class (Example.Semigroup a) => Monoidl a where {\char123}\\ \hspace*{0pt} ~neutral ::~a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}class (Monoidl a) => Monoid a where {\char123}\\ +\hspace*{0pt}class (Example.Monoidl a) => Monoid a where {\char123}\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}class (Monoid a) => Group a where {\char123}\\ +\hspace*{0pt}class (Example.Monoid a) => Group a where {\char123}\\ \hspace*{0pt} ~inverse ::~a -> a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ @@ -1166,32 +1167,32 @@ \hspace*{0pt}inverse{\char95}int ::~Integer -> Integer;\\ \hspace*{0pt}inverse{\char95}int i = negate i;\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Integer where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}int;\\ +\hspace*{0pt}instance Example.Semigroup Integer where {\char123}\\ +\hspace*{0pt} ~mult = Example.mult{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Monoidl Integer where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}int;\\ +\hspace*{0pt}instance Example.Monoidl Integer where {\char123}\\ +\hspace*{0pt} ~neutral = Example.neutral{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Integer where {\char123}\\ +\hspace*{0pt}instance Example.Monoid Integer where {\char123}\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Group Integer where {\char123}\\ -\hspace*{0pt} ~inverse = inverse{\char95}int;\\ +\hspace*{0pt}instance Example.Group Integer where {\char123}\\ +\hspace*{0pt} ~inverse = Example.inverse{\char95}int;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}nat ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow{\char95}nat Zero{\char95}nat x = neutral;\\ -\hspace*{0pt}pow{\char95}nat (Suc n) x = mult x (pow{\char95}nat n x);\\ +\hspace*{0pt}pow{\char95}nat ::~forall a.~(Example.Monoid a) => Example.Nat -> a -> a;\\ +\hspace*{0pt}pow{\char95}nat Example.Zero{\char95}nat x = Example.neutral;\\ +\hspace*{0pt}pow{\char95}nat (Example.Suc n) x = Example.mult x (Example.pow{\char95}nat n x);\\ \hspace*{0pt}\\ -\hspace*{0pt}pow{\char95}int ::~forall a.~(Group a) => Integer -> a -> a;\\ +\hspace*{0pt}pow{\char95}int ::~forall a.~(Example.Group a) => Integer -> a -> a;\\ \hspace*{0pt}pow{\char95}int k x =\\ -\hspace*{0pt} ~(if 0 <= k then pow{\char95}nat (nat k) x\\ -\hspace*{0pt} ~~~else inverse (pow{\char95}nat (nat (negate k)) x));\\ +\hspace*{0pt} ~(if 0 <= k then Example.pow{\char95}nat (Example.nat k) x\\ +\hspace*{0pt} ~~~else Example.inverse (Example.pow{\char95}nat (Example.nat (negate k)) x));\\ \hspace*{0pt}\\ \hspace*{0pt}example ::~Integer;\\ -\hspace*{0pt}example = pow{\char95}int 10 (-2);\\ +\hspace*{0pt}example = Example.pow{\char95}int 10 (-2);\\ \hspace*{0pt}\\ \hspace*{0pt}{\char125}% \end{isamarkuptext}% @@ -1232,15 +1233,15 @@ \hspace*{0pt} ~type 'a group\\ \hspace*{0pt} ~val monoid{\char95}group :~'a group -> 'a monoid\\ \hspace*{0pt} ~val inverse :~'a group -> 'a -> 'a\\ +\hspace*{0pt} ~val pow{\char95}nat :~'a monoid -> nat -> 'a -> 'a\\ +\hspace*{0pt} ~val pow{\char95}int :~'a group -> IntInf.int -> 'a -> 'a\\ \hspace*{0pt} ~val mult{\char95}int :~IntInf.int -> IntInf.int -> IntInf.int\\ +\hspace*{0pt} ~val semigroup{\char95}int :~IntInf.int semigroup\\ \hspace*{0pt} ~val neutral{\char95}int :~IntInf.int\\ -\hspace*{0pt} ~val inverse{\char95}int :~IntInf.int -> IntInf.int\\ -\hspace*{0pt} ~val semigroup{\char95}int :~IntInf.int semigroup\\ \hspace*{0pt} ~val monoidl{\char95}int :~IntInf.int monoidl\\ \hspace*{0pt} ~val monoid{\char95}int :~IntInf.int monoid\\ +\hspace*{0pt} ~val inverse{\char95}int :~IntInf.int -> IntInf.int\\ \hspace*{0pt} ~val group{\char95}int :~IntInf.int group\\ -\hspace*{0pt} ~val pow{\char95}nat :~'a monoid -> nat -> 'a -> 'a\\ -\hspace*{0pt} ~val pow{\char95}int :~'a group -> IntInf.int -> 'a -> 'a\\ \hspace*{0pt} ~val example :~IntInf.int\\ \hspace*{0pt}end = struct\\ \hspace*{0pt}\\ @@ -1266,23 +1267,6 @@ \hspace*{0pt}val monoid{\char95}group = {\char35}monoid{\char95}group :~'a group -> 'a monoid;\\ \hspace*{0pt}val inverse = {\char35}inverse :~'a group -> 'a -> 'a;\\ \hspace*{0pt}\\ -\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ -\hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int);\\ -\hspace*{0pt}\\ -\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoidl{\char95}int =\\ -\hspace*{0pt} ~{\char123}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int monoidl;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val monoid{\char95}int = {\char123}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:~IntInf.int monoid;\\ -\hspace*{0pt}\\ -\hspace*{0pt}val group{\char95}int = {\char123}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ -\hspace*{0pt} ~IntInf.int group;\\ -\hspace*{0pt}\\ \hspace*{0pt}fun pow{\char95}nat A{\char95}~Zero{\char95}nat x = neutral (monoidl{\char95}monoid A{\char95})\\ \hspace*{0pt} ~| pow{\char95}nat A{\char95}~(Suc n) x =\\ \hspace*{0pt} ~~~mult ((semigroup{\char95}monoidl o monoidl{\char95}monoid) A{\char95}) x (pow{\char95}nat A{\char95}~n x);\\ @@ -1292,6 +1276,23 @@ \hspace*{0pt} ~~~then pow{\char95}nat (monoid{\char95}group A{\char95}) (nat k) x\\ \hspace*{0pt} ~~~else inverse A{\char95}~(pow{\char95}nat (monoid{\char95}group A{\char95}) (nat (IntInf.{\char126}~k)) x));\\ \hspace*{0pt}\\ +\hspace*{0pt}fun mult{\char95}int i j = IntInf.+ (i,~j);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val semigroup{\char95}int = {\char123}mult = mult{\char95}int{\char125}~:~IntInf.int semigroup;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val neutral{\char95}int :~IntInf.int = (0 :~IntInf.int);\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoidl{\char95}int =\\ +\hspace*{0pt} ~{\char123}semigroup{\char95}monoidl = semigroup{\char95}int,~neutral = neutral{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int monoidl;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val monoid{\char95}int = {\char123}monoidl{\char95}monoid = monoidl{\char95}int{\char125}~:~IntInf.int monoid;\\ +\hspace*{0pt}\\ +\hspace*{0pt}fun inverse{\char95}int i = IntInf.{\char126}~i;\\ +\hspace*{0pt}\\ +\hspace*{0pt}val group{\char95}int = {\char123}monoid{\char95}group = monoid{\char95}int,~inverse = inverse{\char95}int{\char125}~:\\ +\hspace*{0pt} ~IntInf.int group;\\ +\hspace*{0pt}\\ \hspace*{0pt}val example :~IntInf.int =\\ \hspace*{0pt} ~pow{\char95}int group{\char95}int (10 :~IntInf.int) ({\char126}2 :~IntInf.int);\\ \hspace*{0pt}\\ @@ -1335,9 +1336,6 @@ \noindent% \hspace*{0pt}object Example {\char123}\\ \hspace*{0pt}\\ -\hspace*{0pt}import /*implicits*/ Example.semigroup{\char95}int,~Example.monoidl{\char95}int,\\ -\hspace*{0pt} ~Example.monoid{\char95}int,~Example.group{\char95}int\\ -\hspace*{0pt}\\ \hspace*{0pt}abstract sealed class nat\\ \hspace*{0pt}final case object Zero{\char95}nat extends nat\\ \hspace*{0pt}final case class Suc(a:~nat) extends nat\\ @@ -1348,23 +1346,23 @@ \hspace*{0pt}def nat(i:~BigInt):~nat = nat{\char95}aux(i,~Zero{\char95}nat)\\ \hspace*{0pt}\\ \hspace*{0pt}trait semigroup[A] {\char123}\\ -\hspace*{0pt} ~val `Example+mult`:~(A,~A) => A\\ +\hspace*{0pt} ~val `Example.mult`:~(A,~A) => A\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}def mult[A](a:~A,~b:~A)(implicit A:~semigroup[A]):~A =\\ -\hspace*{0pt} ~A.`Example+mult`(a,~b)\\ +\hspace*{0pt} ~A.`Example.mult`(a,~b)\\ \hspace*{0pt}\\ \hspace*{0pt}trait monoidl[A] extends semigroup[A] {\char123}\\ -\hspace*{0pt} ~val `Example+neutral`:~A\\ +\hspace*{0pt} ~val `Example.neutral`:~A\\ \hspace*{0pt}{\char125}\\ -\hspace*{0pt}def neutral[A](implicit A:~monoidl[A]):~A = A.`Example+neutral`\\ +\hspace*{0pt}def neutral[A](implicit A:~monoidl[A]):~A = A.`Example.neutral`\\ \hspace*{0pt}\\ \hspace*{0pt}trait monoid[A] extends monoidl[A] {\char123}\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}\\ \hspace*{0pt}trait group[A] extends monoid[A] {\char123}\\ -\hspace*{0pt} ~val `Example+inverse`:~A => A\\ +\hspace*{0pt} ~val `Example.inverse`:~A => A\\ \hspace*{0pt}{\char125}\\ -\hspace*{0pt}def inverse[A](a:~A)(implicit A:~group[A]):~A = A.`Example+inverse`(a)\\ +\hspace*{0pt}def inverse[A](a:~A)(implicit A:~group[A]):~A = A.`Example.inverse`(a)\\ \hspace*{0pt}\\ \hspace*{0pt}def pow{\char95}nat[A:~monoid](xa0:~nat,~x:~A):~A = (xa0,~x) match {\char123}\\ \hspace*{0pt} ~case (Zero{\char95}nat,~x) => neutral[A]\\ @@ -1378,27 +1376,27 @@ \hspace*{0pt}def mult{\char95}int(i:~BigInt,~j:~BigInt):~BigInt = i + j\\ \hspace*{0pt}\\ \hspace*{0pt}implicit def semigroup{\char95}int:~semigroup[BigInt] = new semigroup[BigInt] {\char123}\\ -\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ +\hspace*{0pt} ~val `Example.mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}\\ \hspace*{0pt}def neutral{\char95}int:~BigInt = BigInt(0)\\ \hspace*{0pt}\\ \hspace*{0pt}implicit def monoidl{\char95}int:~monoidl[BigInt] = new monoidl[BigInt] {\char123}\\ -\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\ -\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ +\hspace*{0pt} ~val `Example.neutral` = neutral{\char95}int\\ +\hspace*{0pt} ~val `Example.mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}\\ \hspace*{0pt}implicit def monoid{\char95}int:~monoid[BigInt] = new monoid[BigInt] {\char123}\\ -\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\ -\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ +\hspace*{0pt} ~val `Example.neutral` = neutral{\char95}int\\ +\hspace*{0pt} ~val `Example.mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}\\ \hspace*{0pt}def inverse{\char95}int(i:~BigInt):~BigInt = (- i)\\ \hspace*{0pt}\\ \hspace*{0pt}implicit def group{\char95}int:~group[BigInt] = new group[BigInt] {\char123}\\ -\hspace*{0pt} ~val `Example+inverse` = (a:~BigInt) => inverse{\char95}int(a)\\ -\hspace*{0pt} ~val `Example+neutral` = neutral{\char95}int\\ -\hspace*{0pt} ~val `Example+mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ +\hspace*{0pt} ~val `Example.inverse` = (a:~BigInt) => inverse{\char95}int(a)\\ +\hspace*{0pt} ~val `Example.neutral` = neutral{\char95}int\\ +\hspace*{0pt} ~val `Example.mult` = (a:~BigInt,~b:~BigInt) => mult{\char95}int(a,~b)\\ \hspace*{0pt}{\char125}\\ \hspace*{0pt}\\ \hspace*{0pt}def example:~BigInt = pow{\char95}int[BigInt](BigInt(10),~BigInt(- 2))\\
--- a/doc-src/Codegen/Thy/Adaptation.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/Adaptation.thy Fri Sep 03 13:54:39 2010 +0200 @@ -272,12 +272,12 @@ text {* For convenience, the default @{text HOL} setup for @{text Haskell} - maps the @{class eq} class to its counterpart in @{text Haskell}, - giving custom serialisations for the class @{class eq} (by command - @{command_def code_class}) and its operation @{const HOL.eq} + maps the @{class equal} class to its counterpart in @{text Haskell}, + giving custom serialisations for the class @{class equal} (by command + @{command_def code_class}) and its operation @{const HOL.equal} *} -code_class %quotett eq +code_class %quotett equal (Haskell "Eq") code_const %quotett "op =" @@ -285,19 +285,19 @@ text {* \noindent A problem now occurs whenever a type which is an instance - of @{class eq} in @{text HOL} is mapped on a @{text + of @{class equal} in @{text HOL} is mapped on a @{text Haskell}-built-in type which is also an instance of @{text Haskell} @{text Eq}: *} typedecl %quote bar -instantiation %quote bar :: eq +instantiation %quote bar :: equal begin -definition %quote "eq_class.eq (x\<Colon>bar) y \<longleftrightarrow> x = y" +definition %quote "HOL.equal (x\<Colon>bar) y \<longleftrightarrow> x = y" -instance %quote by default (simp add: eq_bar_def) +instance %quote by default (simp add: equal_bar_def) end %quote (*<*) @@ -310,7 +310,7 @@ suppress this additional instance, use @{command_def "code_instance"}: *} -code_instance %quotett bar :: eq +code_instance %quotett bar :: equal (Haskell -)
--- a/doc-src/Codegen/Thy/Evaluation.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/Evaluation.thy Fri Sep 03 13:54:39 2010 +0200 @@ -96,7 +96,7 @@ allows to use pattern matching on constructors stemming from compiled @{text "datatypes"}. - For a less simplistic example, theory @{theory Ferrack} is + For a less simplistic example, theory @{text Ferrack} is a good reference. *}
--- a/doc-src/Codegen/Thy/Further.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/Further.thy Fri Sep 03 13:54:39 2010 +0200 @@ -122,7 +122,7 @@ specific application, you should consider \emph{Imperative Functional Programming with Isabelle/HOL} \cite{bulwahn-et-al:2008:imperative}; the framework described there - is available in session @{theory Imperative_HOL}. + is available in session @{text Imperative_HOL}. *}
--- a/doc-src/Codegen/Thy/Inductive_Predicate.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/Inductive_Predicate.thy Fri Sep 03 13:54:39 2010 +0200 @@ -220,8 +220,8 @@ "values"} and the number of elements. *} -values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 20 "{n. tranclp succ 10 n}" -values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 10 "{n. tranclp succ n 10}" +values %quote [mode: i \<Rightarrow> o \<Rightarrow> bool] 1 "{n. tranclp succ 10 n}" (*FIMXE does not terminate for n\<ge>1*) +values %quote [mode: o \<Rightarrow> i \<Rightarrow> bool] 1 "{n. tranclp succ n 10}" subsection {* Embedding into functional code within Isabelle/HOL *}
--- a/doc-src/Codegen/Thy/Setup.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/Setup.thy Fri Sep 03 13:54:39 2010 +0200 @@ -1,15 +1,12 @@ theory Setup -imports Complex_Main More_List RBT Dlist Mapping +imports + Complex_Main + More_List RBT Dlist Mapping uses "../../antiquote_setup.ML" "../../more_antiquote.ML" begin -ML {* no_document use_thys - ["Efficient_Nat", "Code_Char_chr", "Product_ord", - "~~/src/HOL/Imperative_HOL/Imperative_HOL", - "~~/src/HOL/Decision_Procs/Ferrack"] *} - setup {* let val typ = Simple_Syntax.read_typ;
--- a/doc-src/Codegen/Thy/document/Adaptation.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Adaptation.tex Fri Sep 03 13:54:39 2010 +0200 @@ -499,9 +499,9 @@ % \begin{isamarkuptext}% For convenience, the default \isa{HOL} setup for \isa{Haskell} - maps the \isa{eq} class to its counterpart in \isa{Haskell}, - giving custom serialisations for the class \isa{eq} (by command - \indexdef{}{command}{code\_class}\hypertarget{command.code-class}{\hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}}) and its operation \isa{eq{\isacharunderscore}class{\isachardot}eq}% + maps the \isa{equal} class to its counterpart in \isa{Haskell}, + giving custom serialisations for the class \isa{equal} (by command + \indexdef{}{command}{code\_class}\hypertarget{command.code-class}{\hyperlink{command.code-class}{\mbox{\isa{\isacommand{code{\isacharunderscore}class}}}}}) and its operation \isa{equal{\isacharunderscore}class{\isachardot}equal}% \end{isamarkuptext}% \isamarkuptrue% % @@ -511,7 +511,7 @@ % \isatagquotett \isacommand{code{\isacharunderscore}class}\isamarkupfalse% -\ eq\isanewline +\ equal\isanewline \ \ {\isacharparenleft}Haskell\ {\isachardoublequoteopen}Eq{\isachardoublequoteclose}{\isacharparenright}\isanewline \isanewline \isacommand{code{\isacharunderscore}const}\isamarkupfalse% @@ -526,7 +526,7 @@ % \begin{isamarkuptext}% \noindent A problem now occurs whenever a type which is an instance - of \isa{eq} in \isa{HOL} is mapped on a \isa{Haskell}-built-in type which is also an instance of \isa{Haskell} + of \isa{equal} in \isa{HOL} is mapped on a \isa{Haskell}-built-in type which is also an instance of \isa{Haskell} \isa{Eq}:% \end{isamarkuptext}% \isamarkuptrue% @@ -540,15 +540,15 @@ \ bar\isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ bar\ {\isacharcolon}{\isacharcolon}\ equal\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% -\ {\isachardoublequoteopen}eq{\isacharunderscore}class{\isachardot}eq\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline +\ {\isachardoublequoteopen}HOL{\isachardot}equal\ {\isacharparenleft}x{\isasymColon}bar{\isacharparenright}\ y\ {\isasymlongleftrightarrow}\ x\ {\isacharequal}\ y{\isachardoublequoteclose}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% \ \isacommand{by}\isamarkupfalse% -\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ eq{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline +\ default\ {\isacharparenleft}simp\ add{\isacharcolon}\ equal{\isacharunderscore}bar{\isacharunderscore}def{\isacharparenright}\isanewline \isanewline \isacommand{end}\isamarkupfalse% % @@ -587,7 +587,7 @@ % \isatagquotett \isacommand{code{\isacharunderscore}instance}\isamarkupfalse% -\ bar\ {\isacharcolon}{\isacharcolon}\ eq\isanewline +\ bar\ {\isacharcolon}{\isacharcolon}\ equal\isanewline \ \ {\isacharparenleft}Haskell\ {\isacharminus}{\isacharparenright}% \endisatagquotett {\isafoldquotett}%
--- a/doc-src/Codegen/Thy/document/Evaluation.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Evaluation.tex Fri Sep 03 13:54:39 2010 +0200 @@ -225,7 +225,7 @@ allows to use pattern matching on constructors stemming from compiled \isa{datatypes}. - For a less simplistic example, theory \hyperlink{theory.Ferrack}{\mbox{\isa{Ferrack}}} is + For a less simplistic example, theory \isa{Ferrack} is a good reference.% \end{isamarkuptext}% \isamarkuptrue%
--- a/doc-src/Codegen/Thy/document/Foundations.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Foundations.tex Fri Sep 03 13:54:39 2010 +0200 @@ -247,11 +247,11 @@ \begin{isamarkuptext}% \isatypewriter% \noindent% -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~(if null xs then (Nothing,~AQueue [] [])\\ -\hspace*{0pt} ~~~else dequeue (AQueue [] (reverse xs)));% +\hspace*{0pt}dequeue ::~forall a.~Example.Queue a -> (Maybe a,~Example.Queue a);\\ +\hspace*{0pt}dequeue (Example.AQueue xs (y :~ys)) = (Just y,~Example.AQueue xs ys);\\ +\hspace*{0pt}dequeue (Example.AQueue xs []) =\\ +\hspace*{0pt} ~(if null xs then (Nothing,~Example.AQueue [] [])\\ +\hspace*{0pt} ~~~else Example.dequeue (Example.AQueue [] (reverse xs)));% \end{isamarkuptext}% \isamarkuptrue% % @@ -350,21 +350,21 @@ \isatypewriter% \noindent% \hspace*{0pt}structure Example :~sig\\ -\hspace*{0pt} ~type 'a eq\\ -\hspace*{0pt} ~val eq :~'a eq -> 'a -> 'a -> bool\\ -\hspace*{0pt} ~val eqa :~'a eq -> 'a -> 'a -> bool\\ -\hspace*{0pt} ~val member :~'a eq -> 'a list -> 'a -> bool\\ +\hspace*{0pt} ~type 'a equal\\ +\hspace*{0pt} ~val equal :~'a equal -> 'a -> 'a -> bool\\ +\hspace*{0pt} ~val eq :~'a equal -> 'a -> 'a -> bool\\ +\hspace*{0pt} ~val member :~'a equal -> 'a list -> 'a -> bool\\ \hspace*{0pt} ~val collect{\char95}duplicates :\\ -\hspace*{0pt} ~~~'a eq -> 'a list -> 'a list -> 'a list -> 'a list\\ +\hspace*{0pt} ~~~'a equal -> 'a list -> 'a list -> 'a list -> 'a list\\ \hspace*{0pt}end = struct\\ \hspace*{0pt}\\ -\hspace*{0pt}type 'a eq = {\char123}eq :~'a -> 'a -> bool{\char125};\\ -\hspace*{0pt}val eq = {\char35}eq :~'a eq -> 'a -> 'a -> bool;\\ +\hspace*{0pt}type 'a equal = {\char123}equal :~'a -> 'a -> bool{\char125};\\ +\hspace*{0pt}val equal = {\char35}equal :~'a equal -> 'a -> 'a -> bool;\\ \hspace*{0pt}\\ -\hspace*{0pt}fun eqa A{\char95}~a b = eq A{\char95}~a b;\\ +\hspace*{0pt}fun eq A{\char95}~a b = equal A{\char95}~a b;\\ \hspace*{0pt}\\ \hspace*{0pt}fun member A{\char95}~[] y = false\\ -\hspace*{0pt} ~| member A{\char95}~(x ::~xs) y = eqa A{\char95}~x y orelse member A{\char95}~xs y;\\ +\hspace*{0pt} ~| member A{\char95}~(x ::~xs) y = eq A{\char95}~x y orelse member A{\char95}~xs y;\\ \hspace*{0pt}\\ \hspace*{0pt}fun collect{\char95}duplicates A{\char95}~xs ys [] = xs\\ \hspace*{0pt} ~| collect{\char95}duplicates A{\char95}~xs ys (z ::~zs) =\\ @@ -387,11 +387,11 @@ \begin{isamarkuptext}% \noindent Obviously, polymorphic equality is implemented the Haskell way using a type class. How is this achieved? HOL introduces an - explicit class \isa{eq} with a corresponding operation \isa{eq{\isacharunderscore}class{\isachardot}eq} such that \isa{eq{\isacharunderscore}class{\isachardot}eq\ {\isacharequal}\ op\ {\isacharequal}}. The preprocessing - framework does the rest by propagating the \isa{eq} constraints + explicit class \isa{equal} with a corresponding operation \isa{equal{\isacharunderscore}class{\isachardot}equal} such that \isa{equal{\isacharunderscore}class{\isachardot}equal\ {\isacharequal}\ op\ {\isacharequal}}. The preprocessing + framework does the rest by propagating the \isa{equal} constraints through all dependent code equations. For datatypes, instances of - \isa{eq} are implicitly derived when possible. For other types, - you may instantiate \isa{eq} manually like any other type class.% + \isa{equal} are implicitly derived when possible. For other types, + you may instantiate \isa{equal} manually like any other type class.% \end{isamarkuptext}% \isamarkuptrue% % @@ -444,12 +444,12 @@ \begin{isamarkuptext}% \isatypewriter% \noindent% -\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ +\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Example.Queue a -> (a,~Example.Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue (Example.AQueue xs []) =\\ \hspace*{0pt} ~let {\char123}\\ \hspace*{0pt} ~~~(y :~ys) = reverse xs;\\ -\hspace*{0pt} ~{\char125}~in (y,~AQueue [] ys);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);% +\hspace*{0pt} ~{\char125}~in (y,~Example.AQueue [] ys);\\ +\hspace*{0pt}strict{\char95}dequeue (Example.AQueue xs (y :~ys)) = (y,~Example.AQueue xs ys);% \end{isamarkuptext}% \isamarkuptrue% % @@ -538,11 +538,11 @@ \hspace*{0pt}empty{\char95}queue ::~forall a.~a;\\ \hspace*{0pt}empty{\char95}queue = error {\char34}empty{\char95}queue{\char34};\\ \hspace*{0pt}\\ -\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Queue a -> (a,~Queue a);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs (y :~ys)) = (y,~AQueue xs ys);\\ -\hspace*{0pt}strict{\char95}dequeue (AQueue xs []) =\\ -\hspace*{0pt} ~(if null xs then empty{\char95}queue\\ -\hspace*{0pt} ~~~else strict{\char95}dequeue (AQueue [] (reverse xs)));% +\hspace*{0pt}strict{\char95}dequeue ::~forall a.~Example.Queue a -> (a,~Example.Queue a);\\ +\hspace*{0pt}strict{\char95}dequeue (Example.AQueue xs (y :~ys)) = (y,~Example.AQueue xs ys);\\ +\hspace*{0pt}strict{\char95}dequeue (Example.AQueue xs []) =\\ +\hspace*{0pt} ~(if null xs then Example.empty{\char95}queue\\ +\hspace*{0pt} ~~~else Example.strict{\char95}dequeue (Example.AQueue [] (reverse xs)));% \end{isamarkuptext}% \isamarkuptrue% %
--- a/doc-src/Codegen/Thy/document/Further.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Further.tex Fri Sep 03 13:54:39 2010 +0200 @@ -216,13 +216,13 @@ \begin{isamarkuptext}% \isatypewriter% \noindent% -\hspace*{0pt}funpow ::~forall a.~Nat -> (a -> a) -> a -> a;\\ -\hspace*{0pt}funpow Zero{\char95}nat f = id;\\ -\hspace*{0pt}funpow (Suc n) f = f .~funpow n f;\\ +\hspace*{0pt}funpow ::~forall a.~Example.Nat -> (a -> a) -> a -> a;\\ +\hspace*{0pt}funpow Example.Zero{\char95}nat f = id;\\ +\hspace*{0pt}funpow (Example.Suc n) f = f .~Example.funpow n f;\\ \hspace*{0pt}\\ -\hspace*{0pt}funpows ::~forall a.~[Nat] -> (a -> a) -> a -> a;\\ +\hspace*{0pt}funpows ::~forall a.~[Example.Nat] -> (a -> a) -> a -> a;\\ \hspace*{0pt}funpows [] = id;\\ -\hspace*{0pt}funpows (x :~xs) = funpow x .~funpows xs;% +\hspace*{0pt}funpows (x :~xs) = Example.funpow x .~Example.funpows xs;% \end{isamarkuptext}% \isamarkuptrue% % @@ -242,7 +242,7 @@ specific application, you should consider \emph{Imperative Functional Programming with Isabelle/HOL} \cite{bulwahn-et-al:2008:imperative}; the framework described there - is available in session \hyperlink{theory.Imperative-HOL}{\mbox{\isa{Imperative{\isacharunderscore}HOL}}}.% + is available in session \isa{Imperative{\isacharunderscore}HOL}.% \end{isamarkuptext}% \isamarkuptrue% %
--- a/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Inductive_Predicate.tex Fri Sep 03 13:54:39 2010 +0200 @@ -422,9 +422,9 @@ % \isatagquote \isacommand{values}\isamarkupfalse% -\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{2}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\isanewline +\ {\isacharbrackleft}mode{\isacharcolon}\ i\ {\isasymRightarrow}\ o\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{1}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ {\isadigit{1}}{\isadigit{0}}\ n{\isacharbraceright}{\isachardoublequoteclose}\ \isanewline \isacommand{values}\isamarkupfalse% -\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{1}}{\isadigit{0}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}% +\ {\isacharbrackleft}mode{\isacharcolon}\ o\ {\isasymRightarrow}\ i\ {\isasymRightarrow}\ bool{\isacharbrackright}\ {\isadigit{1}}\ {\isachardoublequoteopen}{\isacharbraceleft}n{\isachardot}\ tranclp\ succ\ n\ {\isadigit{1}}{\isadigit{0}}{\isacharbraceright}{\isachardoublequoteclose}% \endisatagquote {\isafoldquote}% %
--- a/doc-src/Codegen/Thy/document/Introduction.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Introduction.tex Fri Sep 03 13:54:39 2010 +0200 @@ -231,19 +231,19 @@ \hspace*{0pt}\\ \hspace*{0pt}data Queue a = AQueue [a] [a];\\ \hspace*{0pt}\\ -\hspace*{0pt}empty ::~forall a.~Queue a;\\ -\hspace*{0pt}empty = AQueue [] [];\\ +\hspace*{0pt}empty ::~forall a.~Example.Queue a;\\ +\hspace*{0pt}empty = Example.AQueue [] [];\\ \hspace*{0pt}\\ -\hspace*{0pt}dequeue ::~forall a.~Queue a -> (Maybe a,~Queue a);\\ -\hspace*{0pt}dequeue (AQueue [] []) = (Nothing,~AQueue [] []);\\ -\hspace*{0pt}dequeue (AQueue xs (y :~ys)) = (Just y,~AQueue xs ys);\\ -\hspace*{0pt}dequeue (AQueue (v :~va) []) =\\ +\hspace*{0pt}dequeue ::~forall a.~Example.Queue a -> (Maybe a,~Example.Queue a);\\ +\hspace*{0pt}dequeue (Example.AQueue [] []) = (Nothing,~Example.AQueue [] []);\\ +\hspace*{0pt}dequeue (Example.AQueue xs (y :~ys)) = (Just y,~Example.AQueue xs ys);\\ +\hspace*{0pt}dequeue (Example.AQueue (v :~va) []) =\\ \hspace*{0pt} ~let {\char123}\\ \hspace*{0pt} ~~~(y :~ys) = reverse (v :~va);\\ -\hspace*{0pt} ~{\char125}~in (Just y,~AQueue [] ys);\\ +\hspace*{0pt} ~{\char125}~in (Just y,~Example.AQueue [] ys);\\ \hspace*{0pt}\\ -\hspace*{0pt}enqueue ::~forall a.~a -> Queue a -> Queue a;\\ -\hspace*{0pt}enqueue x (AQueue xs ys) = AQueue (x :~xs) ys;\\ +\hspace*{0pt}enqueue ::~forall a.~a -> Example.Queue a -> Example.Queue a;\\ +\hspace*{0pt}enqueue x (Example.AQueue xs ys) = Example.AQueue (x :~xs) ys;\\ \hspace*{0pt}\\ \hspace*{0pt}{\char125}% \end{isamarkuptext}% @@ -397,41 +397,41 @@ \noindent% \hspace*{0pt}module Example where {\char123}\\ \hspace*{0pt}\\ -\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Nat;\\ +\hspace*{0pt}data Nat = Zero{\char95}nat | Suc Example.Nat;\\ \hspace*{0pt}\\ -\hspace*{0pt}plus{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}plus{\char95}nat (Suc m) n = plus{\char95}nat m (Suc n);\\ -\hspace*{0pt}plus{\char95}nat Zero{\char95}nat n = n;\\ +\hspace*{0pt}plus{\char95}nat ::~Example.Nat -> Example.Nat -> Example.Nat;\\ +\hspace*{0pt}plus{\char95}nat (Example.Suc m) n = Example.plus{\char95}nat m (Example.Suc n);\\ +\hspace*{0pt}plus{\char95}nat Example.Zero{\char95}nat n = n;\\ \hspace*{0pt}\\ \hspace*{0pt}class Semigroup a where {\char123}\\ \hspace*{0pt} ~mult ::~a -> a -> a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}class (Semigroup a) => Monoid a where {\char123}\\ +\hspace*{0pt}class (Example.Semigroup a) => Monoid a where {\char123}\\ \hspace*{0pt} ~neutral ::~a;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}pow ::~forall a.~(Monoid a) => Nat -> a -> a;\\ -\hspace*{0pt}pow Zero{\char95}nat a = neutral;\\ -\hspace*{0pt}pow (Suc n) a = mult a (pow n a);\\ +\hspace*{0pt}pow ::~forall a.~(Example.Monoid a) => Example.Nat -> a -> a;\\ +\hspace*{0pt}pow Example.Zero{\char95}nat a = Example.neutral;\\ +\hspace*{0pt}pow (Example.Suc n) a = Example.mult a (Example.pow n a);\\ \hspace*{0pt}\\ -\hspace*{0pt}mult{\char95}nat ::~Nat -> Nat -> Nat;\\ -\hspace*{0pt}mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat;\\ -\hspace*{0pt}mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ +\hspace*{0pt}mult{\char95}nat ::~Example.Nat -> Example.Nat -> Example.Nat;\\ +\hspace*{0pt}mult{\char95}nat Example.Zero{\char95}nat n = Example.Zero{\char95}nat;\\ +\hspace*{0pt}mult{\char95}nat (Example.Suc m) n = Example.plus{\char95}nat n (Example.mult{\char95}nat m n);\\ \hspace*{0pt}\\ -\hspace*{0pt}neutral{\char95}nat ::~Nat;\\ -\hspace*{0pt}neutral{\char95}nat = Suc Zero{\char95}nat;\\ +\hspace*{0pt}neutral{\char95}nat ::~Example.Nat;\\ +\hspace*{0pt}neutral{\char95}nat = Example.Suc Example.Zero{\char95}nat;\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Semigroup Nat where {\char123}\\ -\hspace*{0pt} ~mult = mult{\char95}nat;\\ +\hspace*{0pt}instance Example.Semigroup Example.Nat where {\char123}\\ +\hspace*{0pt} ~mult = Example.mult{\char95}nat;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}instance Monoid Nat where {\char123}\\ -\hspace*{0pt} ~neutral = neutral{\char95}nat;\\ +\hspace*{0pt}instance Example.Monoid Example.Nat where {\char123}\\ +\hspace*{0pt} ~neutral = Example.neutral{\char95}nat;\\ \hspace*{0pt}{\char125};\\ \hspace*{0pt}\\ -\hspace*{0pt}bexp ::~Nat -> Nat;\\ -\hspace*{0pt}bexp n = pow n (Suc (Suc Zero{\char95}nat));\\ +\hspace*{0pt}bexp ::~Example.Nat -> Example.Nat;\\ +\hspace*{0pt}bexp n = Example.pow n (Example.Suc (Example.Suc Example.Zero{\char95}nat));\\ \hspace*{0pt}\\ \hspace*{0pt}{\char125}% \end{isamarkuptext}% @@ -470,8 +470,8 @@ \hspace*{0pt} ~val neutral :~'a monoid -> 'a\\ \hspace*{0pt} ~val pow :~'a monoid -> nat -> 'a -> 'a\\ \hspace*{0pt} ~val mult{\char95}nat :~nat -> nat -> nat\\ +\hspace*{0pt} ~val semigroup{\char95}nat :~nat semigroup\\ \hspace*{0pt} ~val neutral{\char95}nat :~nat\\ -\hspace*{0pt} ~val semigroup{\char95}nat :~nat semigroup\\ \hspace*{0pt} ~val monoid{\char95}nat :~nat monoid\\ \hspace*{0pt} ~val bexp :~nat -> nat\\ \hspace*{0pt}end = struct\\ @@ -494,9 +494,9 @@ \hspace*{0pt}fun mult{\char95}nat Zero{\char95}nat n = Zero{\char95}nat\\ \hspace*{0pt} ~| mult{\char95}nat (Suc m) n = plus{\char95}nat n (mult{\char95}nat m n);\\ \hspace*{0pt}\\ -\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat;\\ +\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ \hspace*{0pt}\\ -\hspace*{0pt}val semigroup{\char95}nat = {\char123}mult = mult{\char95}nat{\char125}~:~nat semigroup;\\ +\hspace*{0pt}val neutral{\char95}nat :~nat = Suc Zero{\char95}nat;\\ \hspace*{0pt}\\ \hspace*{0pt}val monoid{\char95}nat = {\char123}semigroup{\char95}monoid = semigroup{\char95}nat,~neutral = neutral{\char95}nat{\char125}\\ \hspace*{0pt} ~:~nat monoid;\\
--- a/doc-src/Codegen/Thy/document/Refinement.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/document/Refinement.tex Fri Sep 03 13:54:39 2010 +0200 @@ -74,10 +74,11 @@ \begin{isamarkuptext}% \isatypewriter% \noindent% -\hspace*{0pt}fib ::~Nat -> Nat;\\ -\hspace*{0pt}fib Zero{\char95}nat = Zero{\char95}nat;\\ -\hspace*{0pt}fib (Suc Zero{\char95}nat) = Suc Zero{\char95}nat;\\ -\hspace*{0pt}fib (Suc (Suc n)) = plus{\char95}nat (fib n) (fib (Suc n));% +\hspace*{0pt}fib ::~Example.Nat -> Example.Nat;\\ +\hspace*{0pt}fib Example.Zero{\char95}nat = Example.Zero{\char95}nat;\\ +\hspace*{0pt}fib (Example.Suc Example.Zero{\char95}nat) = Example.Suc Example.Zero{\char95}nat;\\ +\hspace*{0pt}fib (Example.Suc (Example.Suc n)) =\\ +\hspace*{0pt} ~Example.plus{\char95}nat (Example.fib n) (Example.fib (Example.Suc n));% \end{isamarkuptext}% \isamarkuptrue% % @@ -172,15 +173,17 @@ \begin{isamarkuptext}% \isatypewriter% \noindent% -\hspace*{0pt}fib{\char95}step ::~Nat -> (Nat,~Nat);\\ -\hspace*{0pt}fib{\char95}step (Suc n) = let {\char123}\\ -\hspace*{0pt} ~~~~~~~~~~~~~~~~~~~~(m,~q) = fib{\char95}step n;\\ -\hspace*{0pt} ~~~~~~~~~~~~~~~~~~{\char125}~in (plus{\char95}nat m q,~m);\\ -\hspace*{0pt}fib{\char95}step Zero{\char95}nat = (Suc Zero{\char95}nat,~Zero{\char95}nat);\\ +\hspace*{0pt}fib{\char95}step ::~Example.Nat -> (Example.Nat,~Example.Nat);\\ +\hspace*{0pt}fib{\char95}step (Example.Suc n) =\\ +\hspace*{0pt} ~let {\char123}\\ +\hspace*{0pt} ~~~(m,~q) = Example.fib{\char95}step n;\\ +\hspace*{0pt} ~{\char125}~in (Example.plus{\char95}nat m q,~m);\\ +\hspace*{0pt}fib{\char95}step Example.Zero{\char95}nat =\\ +\hspace*{0pt} ~(Example.Suc Example.Zero{\char95}nat,~Example.Zero{\char95}nat);\\ \hspace*{0pt}\\ -\hspace*{0pt}fib ::~Nat -> Nat;\\ -\hspace*{0pt}fib (Suc n) = fst (fib{\char95}step n);\\ -\hspace*{0pt}fib Zero{\char95}nat = Zero{\char95}nat;% +\hspace*{0pt}fib ::~Example.Nat -> Example.Nat;\\ +\hspace*{0pt}fib (Example.Suc n) = fst (Example.fib{\char95}step n);\\ +\hspace*{0pt}fib Example.Zero{\char95}nat = Example.Zero{\char95}nat;% \end{isamarkuptext}% \isamarkuptrue% % @@ -590,28 +593,30 @@ \hspace*{0pt}\\ \hspace*{0pt}newtype Dlist a = Dlist [a];\\ \hspace*{0pt}\\ -\hspace*{0pt}empty ::~forall a.~Dlist a;\\ -\hspace*{0pt}empty = Dlist [];\\ +\hspace*{0pt}empty ::~forall a.~Example.Dlist a;\\ +\hspace*{0pt}empty = Example.Dlist [];\\ \hspace*{0pt}\\ \hspace*{0pt}member ::~forall a.~(Eq a) => [a] -> a -> Bool;\\ \hspace*{0pt}member [] y = False;\\ -\hspace*{0pt}member (x :~xs) y = x == y || member xs y;\\ +\hspace*{0pt}member (x :~xs) y = x == y || Example.member xs y;\\ \hspace*{0pt}\\ \hspace*{0pt}inserta ::~forall a.~(Eq a) => a -> [a] -> [a];\\ -\hspace*{0pt}inserta x xs = (if member xs x then xs else x :~xs);\\ +\hspace*{0pt}inserta x xs = (if Example.member xs x then xs else x :~xs);\\ \hspace*{0pt}\\ -\hspace*{0pt}list{\char95}of{\char95}dlist ::~forall a.~Dlist a -> [a];\\ -\hspace*{0pt}list{\char95}of{\char95}dlist (Dlist x) = x;\\ +\hspace*{0pt}list{\char95}of{\char95}dlist ::~forall a.~Example.Dlist a -> [a];\\ +\hspace*{0pt}list{\char95}of{\char95}dlist (Example.Dlist x) = x;\\ \hspace*{0pt}\\ -\hspace*{0pt}insert ::~forall a.~(Eq a) => a -> Dlist a -> Dlist a;\\ -\hspace*{0pt}insert x dxs = Dlist (inserta x (list{\char95}of{\char95}dlist dxs));\\ +\hspace*{0pt}insert ::~forall a.~(Eq a) => a -> Example.Dlist a -> Example.Dlist a;\\ +\hspace*{0pt}insert x dxs =\\ +\hspace*{0pt} ~Example.Dlist (Example.inserta x (Example.list{\char95}of{\char95}dlist dxs));\\ \hspace*{0pt}\\ \hspace*{0pt}remove1 ::~forall a.~(Eq a) => a -> [a] -> [a];\\ \hspace*{0pt}remove1 x [] = [];\\ -\hspace*{0pt}remove1 x (y :~xs) = (if x == y then xs else y :~remove1 x xs);\\ +\hspace*{0pt}remove1 x (y :~xs) = (if x == y then xs else y :~Example.remove1 x xs);\\ \hspace*{0pt}\\ -\hspace*{0pt}remove ::~forall a.~(Eq a) => a -> Dlist a -> Dlist a;\\ -\hspace*{0pt}remove x dxs = Dlist (remove1 x (list{\char95}of{\char95}dlist dxs));\\ +\hspace*{0pt}remove ::~forall a.~(Eq a) => a -> Example.Dlist a -> Example.Dlist a;\\ +\hspace*{0pt}remove x dxs =\\ +\hspace*{0pt} ~Example.Dlist (Example.remove1 x (Example.list{\char95}of{\char95}dlist dxs));\\ \hspace*{0pt}\\ \hspace*{0pt}{\char125}% \end{isamarkuptext}%
--- a/doc-src/Codegen/Thy/examples/Example.hs Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/Codegen/Thy/examples/Example.hs Fri Sep 03 13:54:39 2010 +0200 @@ -4,18 +4,18 @@ data Queue a = AQueue [a] [a]; -empty :: forall a. Queue a; -empty = AQueue [] []; +empty :: forall a. Example.Queue a; +empty = Example.AQueue [] []; -dequeue :: forall a. Queue a -> (Maybe a, Queue a); -dequeue (AQueue [] []) = (Nothing, AQueue [] []); -dequeue (AQueue xs (y : ys)) = (Just y, AQueue xs ys); -dequeue (AQueue (v : va) []) = +dequeue :: forall a. Example.Queue a -> (Maybe a, Example.Queue a); +dequeue (Example.AQueue [] []) = (Nothing, Example.AQueue [] []); +dequeue (Example.AQueue xs (y : ys)) = (Just y, Example.AQueue xs ys); +dequeue (Example.AQueue (v : va) []) = let { (y : ys) = reverse (v : va); - } in (Just y, AQueue [] ys); + } in (Just y, Example.AQueue [] ys); -enqueue :: forall a. a -> Queue a -> Queue a; -enqueue x (AQueue xs ys) = AQueue (x : xs) ys; +enqueue :: forall a. a -> Example.Queue a -> Example.Queue a; +enqueue x (Example.AQueue xs ys) = Example.AQueue (x : xs) ys; }
--- a/doc-src/IsarRef/Thy/Inner_Syntax.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy Fri Sep 03 13:54:39 2010 +0200 @@ -98,7 +98,7 @@ \begin{mldecls} @{index_ML show_types: "bool Unsynchronized.ref"} & default @{ML false} \\ @{index_ML show_sorts: "bool Unsynchronized.ref"} & default @{ML false} \\ - @{index_ML show_consts: "bool Unsynchronized.ref"} & default @{ML false} \\ + @{index_ML show_consts: "bool Config.T"} & default @{ML false} \\ @{index_ML long_names: "bool Unsynchronized.ref"} & default @{ML false} \\ @{index_ML short_names: "bool Unsynchronized.ref"} & default @{ML false} \\ @{index_ML unique_names: "bool Unsynchronized.ref"} & default @{ML true} \\
--- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Fri Sep 03 13:54:04 2010 +0200 +++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Fri Sep 03 13:54:39 2010 +0200 @@ -120,7 +120,7 @@ \begin{mldecls} \indexdef{}{ML}{show\_types}\verb|show_types: bool Unsynchronized.ref| & default \verb|false| \\ \indexdef{}{ML}{show\_sorts}\verb|show_sorts: bool Unsynchronized.ref| & default \verb|false| \\ - \indexdef{}{ML}{show\_consts}\verb|show_consts: bool Unsynchronized.ref| & default \verb|false| \\ + \indexdef{}{ML}{show\_consts}\verb|show_consts: bool Config.T| & default \verb|false| \\ \indexdef{}{ML}{long\_names}\verb|long_names: bool Unsynchronized.ref| & default \verb|false| \\ \indexdef{}{ML}{short\_names}\verb|short_names: bool Unsynchronized.ref| & default \verb|false| \\ \indexdef{}{ML}{unique\_names}\verb|unique_names: bool Unsynchronized.ref| & default \verb|true| \\
--- a/src/HOL/Fun.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Fun.thy Fri Sep 03 13:54:39 2010 +0200 @@ -117,31 +117,27 @@ no_notation fcomp (infixl "\<circ>>" 60) -subsection {* Injectivity and Surjectivity *} +subsection {* Injectivity, Surjectivity and Bijectivity *} + +definition inj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool" where -- "injective" + "inj_on f A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. f x = f y \<longrightarrow> x = y)" -definition - inj_on :: "['a => 'b, 'a set] => bool" where - -- "injective" - "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" +definition surj_on :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b set \<Rightarrow> bool" where -- "surjective" + "surj_on f B \<longleftrightarrow> B \<subseteq> range f" + +definition bij_betw :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> bool" where -- "bijective" + "bij_betw f A B \<longleftrightarrow> inj_on f A \<and> f ` A = B" text{*A common special case: functions injective over the entire domain type.*} abbreviation - "inj f == inj_on f UNIV" - -definition - bij_betw :: "('a => 'b) => 'a set => 'b set => bool" where -- "bijective" - "bij_betw f A B \<longleftrightarrow> inj_on f A & f ` A = B" + "inj f \<equiv> inj_on f UNIV" -definition - surj :: "('a => 'b) => bool" where - -- "surjective" - "surj f == ! y. ? x. y=f(x)" +abbreviation + "surj f \<equiv> surj_on f UNIV" -definition - bij :: "('a => 'b) => bool" where - -- "bijective" - "bij f == inj f & surj f" +abbreviation + "bij f \<equiv> bij_betw f UNIV UNIV" lemma injI: assumes "\<And>x y. f x = f y \<Longrightarrow> x = y" @@ -173,16 +169,16 @@ by (simp add: inj_on_eq_iff) lemma inj_on_id[simp]: "inj_on id A" - by (simp add: inj_on_def) + by (simp add: inj_on_def) lemma inj_on_id2[simp]: "inj_on (%x. x) A" -by (simp add: inj_on_def) +by (simp add: inj_on_def) -lemma surj_id[simp]: "surj id" -by (simp add: surj_def) +lemma surj_id[simp]: "surj_on id A" +by (simp add: surj_on_def) lemma bij_id[simp]: "bij id" -by (simp add: bij_def) +by (simp add: bij_betw_def) lemma inj_onI: "(!! x y. [| x:A; y:A; f(x) = f(y) |] ==> x=y) ==> inj_on f A" @@ -242,19 +238,26 @@ apply (blast) done -lemma surjI: "(!! x. g(f x) = x) ==> surj g" -apply (simp add: surj_def) -apply (blast intro: sym) -done +lemma surj_onI: "(\<And>x. x \<in> B \<Longrightarrow> g (f x) = x) \<Longrightarrow> surj_on g B" + by (simp add: surj_on_def) (blast intro: sym) + +lemma surj_onD: "surj_on f B \<Longrightarrow> y \<in> B \<Longrightarrow> \<exists>x. y = f x" + by (auto simp: surj_on_def) + +lemma surj_on_range_iff: "surj_on f B \<longleftrightarrow> (\<exists>A. f ` A = B)" + unfolding surj_on_def by (auto intro!: exI[of _ "f -` B"]) -lemma surj_range: "surj f ==> range f = UNIV" -by (auto simp add: surj_def) +lemma surj_def: "surj f \<longleftrightarrow> (\<forall>y. \<exists>x. y = f x)" + by (simp add: surj_on_def subset_eq image_iff) + +lemma surjI: "(\<And> x. g (f x) = x) \<Longrightarrow> surj g" + by (blast intro: surj_onI) -lemma surjD: "surj f ==> EX x. y = f x" -by (simp add: surj_def) +lemma surjD: "surj f \<Longrightarrow> \<exists>x. y = f x" + by (simp add: surj_def) -lemma surjE: "surj f ==> (!!x. y = f x ==> C) ==> C" -by (simp add: surj_def, blast) +lemma surjE: "surj f \<Longrightarrow> (\<And>x. y = f x \<Longrightarrow> C) \<Longrightarrow> C" + by (simp add: surj_def, blast) lemma comp_surj: "[| surj f; surj g |] ==> surj (g o f)" apply (simp add: comp_def surj_def, clarify) @@ -262,6 +265,18 @@ apply (drule_tac x = x in spec, blast) done +lemma surj_range: "surj f \<Longrightarrow> range f = UNIV" + by (auto simp add: surj_on_def) + +lemma surj_range_iff: "surj f \<longleftrightarrow> range f = UNIV" + unfolding surj_on_def by auto + +lemma bij_betw_imp_surj: "bij_betw f A UNIV \<Longrightarrow> surj f" + unfolding bij_betw_def surj_range_iff by auto + +lemma bij_def: "bij f \<longleftrightarrow> inj f \<and> surj f" + unfolding surj_range_iff bij_betw_def .. + lemma bijI: "[| inj f; surj f |] ==> bij f" by (simp add: bij_def) @@ -274,6 +289,9 @@ lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A" by (simp add: bij_betw_def) +lemma bij_betw_imp_surj_on: "bij_betw f A B \<Longrightarrow> surj_on f B" +by (auto simp: bij_betw_def surj_on_range_iff) + lemma bij_comp: "bij f \<Longrightarrow> bij g \<Longrightarrow> bij (g o f)" by(fastsimp intro: comp_inj_on comp_surj simp: bij_def surj_range) @@ -312,6 +330,11 @@ ultimately show ?thesis by(auto simp:bij_betw_def) qed +lemma bij_betw_combine: + assumes "bij_betw f A B" "bij_betw f C D" "B \<inter> D = {}" + shows "bij_betw f (A \<union> C) (B \<union> D)" + using assms unfolding bij_betw_def inj_on_Un image_Un by auto + lemma surj_image_vimage_eq: "surj f ==> f ` (f -` A) = A" by (simp add: surj_range) @@ -497,44 +520,46 @@ lemma comp_swap: "f \<circ> swap a b g = swap a b (f \<circ> g)" by (rule ext, simp add: fun_upd_def swap_def) +lemma swap_image_eq [simp]: + assumes "a \<in> A" "b \<in> A" shows "swap a b f ` A = f ` A" +proof - + have subset: "\<And>f. swap a b f ` A \<subseteq> f ` A" + using assms by (auto simp: image_iff swap_def) + then have "swap a b (swap a b f) ` A \<subseteq> (swap a b f) ` A" . + with subset[of f] show ?thesis by auto +qed + lemma inj_on_imp_inj_on_swap: - "[|inj_on f A; a \<in> A; b \<in> A|] ==> inj_on (swap a b f) A" -by (simp add: inj_on_def swap_def, blast) + "\<lbrakk>inj_on f A; a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> inj_on (swap a b f) A" + by (simp add: inj_on_def swap_def, blast) lemma inj_on_swap_iff [simp]: - assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A = inj_on f A" -proof + assumes A: "a \<in> A" "b \<in> A" shows "inj_on (swap a b f) A \<longleftrightarrow> inj_on f A" +proof assume "inj_on (swap a b f) A" - with A have "inj_on (swap a b (swap a b f)) A" - by (iprover intro: inj_on_imp_inj_on_swap) - thus "inj_on f A" by simp + with A have "inj_on (swap a b (swap a b f)) A" + by (iprover intro: inj_on_imp_inj_on_swap) + thus "inj_on f A" by simp next assume "inj_on f A" with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) qed -lemma surj_imp_surj_swap: "surj f ==> surj (swap a b f)" -apply (simp add: surj_def swap_def, clarify) -apply (case_tac "y = f b", blast) -apply (case_tac "y = f a", auto) -done +lemma surj_imp_surj_swap: "surj f \<Longrightarrow> surj (swap a b f)" + unfolding surj_range_iff by simp + +lemma surj_swap_iff [simp]: "surj (swap a b f) \<longleftrightarrow> surj f" + unfolding surj_range_iff by simp -lemma surj_swap_iff [simp]: "surj (swap a b f) = surj f" -proof - assume "surj (swap a b f)" - hence "surj (swap a b (swap a b f))" by (rule surj_imp_surj_swap) - thus "surj f" by simp -next - assume "surj f" - thus "surj (swap a b f)" by (rule surj_imp_surj_swap) -qed +lemma bij_betw_swap_iff [simp]: + "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> bij_betw (swap x y f) A B \<longleftrightarrow> bij_betw f A B" + by (auto simp: bij_betw_def) -lemma bij_swap_iff: "bij (swap a b f) = bij f" -by (simp add: bij_def) +lemma bij_swap_iff [simp]: "bij (swap a b f) \<longleftrightarrow> bij f" + by simp hide_const (open) swap - subsection {* Inversion of injective functions *} definition the_inv_into :: "'a set => ('a => 'b) => ('b => 'a)" where
--- a/src/HOL/IsaMakefile Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/IsaMakefile Fri Sep 03 13:54:39 2010 +0200 @@ -209,7 +209,6 @@ Tools/primrec.ML \ Tools/prop_logic.ML \ Tools/refute.ML \ - Tools/refute_isar.ML \ Tools/rewrite_hol_proof.ML \ Tools/sat_funcs.ML \ Tools/sat_solver.ML \
--- a/src/HOL/Library/Permutation.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Library/Permutation.thy Fri Sep 03 13:54:39 2010 +0200 @@ -183,4 +183,55 @@ lemma perm_remdups_iff_eq_set: "remdups x <~~> remdups y = (set x = set y)" by (metis List.set_remdups perm_set_eq eq_set_perm_remdups) +lemma permutation_Ex_bij: + assumes "xs <~~> ys" + shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))" +using assms proof induct + case Nil then show ?case unfolding bij_betw_def by simp +next + case (swap y x l) + show ?case + proof (intro exI[of _ "Fun.swap 0 1 id"] conjI allI impI) + show "bij_betw (Fun.swap 0 1 id) {..<length (y # x # l)} {..<length (x # y # l)}" + by (auto simp: bij_betw_def bij_betw_swap_iff) + fix i assume "i < length(y#x#l)" + show "(y # x # l) ! i = (x # y # l) ! (Fun.swap 0 1 id) i" + by (cases i) (auto simp: Fun.swap_def gr0_conv_Suc) + qed +next + case (Cons xs ys z) + then obtain f where bij: "bij_betw f {..<length xs} {..<length ys}" and + perm: "\<forall>i<length xs. xs ! i = ys ! (f i)" by blast + let "?f i" = "case i of Suc n \<Rightarrow> Suc (f n) | 0 \<Rightarrow> 0" + show ?case + proof (intro exI[of _ ?f] allI conjI impI) + have *: "{..<length (z#xs)} = {0} \<union> Suc ` {..<length xs}" + "{..<length (z#ys)} = {0} \<union> Suc ` {..<length ys}" + by (simp_all add: lessThan_Suc_eq_insert_0) + show "bij_betw ?f {..<length (z#xs)} {..<length (z#ys)}" unfolding * + proof (rule bij_betw_combine) + show "bij_betw ?f (Suc ` {..<length xs}) (Suc ` {..<length ys})" + using bij unfolding bij_betw_def + by (auto intro!: inj_onI imageI dest: inj_onD simp: image_compose[symmetric] comp_def) + qed (auto simp: bij_betw_def) + fix i assume "i < length (z#xs)" + then show "(z # xs) ! i = (z # ys) ! (?f i)" + using perm by (cases i) auto + qed +next + case (trans xs ys zs) + then obtain f g where + bij: "bij_betw f {..<length xs} {..<length ys}" "bij_betw g {..<length ys} {..<length zs}" and + perm: "\<forall>i<length xs. xs ! i = ys ! (f i)" "\<forall>i<length ys. ys ! i = zs ! (g i)" by blast + show ?case + proof (intro exI[of _ "g\<circ>f"] conjI allI impI) + show "bij_betw (g \<circ> f) {..<length xs} {..<length zs}" + using bij by (rule bij_betw_trans) + fix i assume "i < length xs" + with bij have "f i < length ys" unfolding bij_betw_def by force + with `i < length xs` show "xs ! i = zs ! (g \<circ> f) i" + using trans(1,3)[THEN perm_length] perm by force + qed +qed + end
--- a/src/HOL/List.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/List.thy Fri Sep 03 13:54:39 2010 +0200 @@ -3076,6 +3076,10 @@ "\<not> P x \<Longrightarrow> remove1 x (filter P xs) = filter P xs" by(induct xs) auto +lemma filter_remove1: + "filter Q (remove1 x xs) = remove1 x (filter Q xs)" +by (induct xs) auto + lemma notin_set_remove1[simp]: "x ~: set xs ==> x ~: set(remove1 y xs)" apply(insert set_remove1_subset) apply fast
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Sep 03 13:54:39 2010 +0200 @@ -5027,7 +5027,7 @@ (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i<DIM('a). ((a$$i \<le> x$$i \<and> x$$i \<le> b$$i) \<or> (b$$i \<le> x$$i \<and> x$$i \<le> a$$i))) \<longrightarrow> x \<in> s)" lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1) - "is_interval {a<..<b}" (is ?th2) proof - + "is_interval {a<..<b}" (is ?th2) proof - have *:"\<And>x y z::real. x < y \<Longrightarrow> y < z \<Longrightarrow> x < z" by auto show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff by(meson order_trans le_less_trans less_le_trans *)+ qed @@ -5051,6 +5051,9 @@ lemma continuous_at_inner: "continuous (at x) (inner a)" unfolding continuous_at by (intro tendsto_intros) +lemma continuous_at_euclidean_component[intro!, simp]: "continuous (at x) (\<lambda>x. x $$ i)" + unfolding euclidean_component_def by (rule continuous_at_inner) + lemma continuous_on_inner: fixes s :: "'a::real_inner set" shows "continuous_on s (inner a)" @@ -5159,6 +5162,9 @@ by (simp add: closed_def open_halfspace_component_lt) qed +lemma open_vimage_euclidean_component: "open S \<Longrightarrow> open ((\<lambda>x. x $$ i) -` S)" + by (auto intro!: continuous_open_vimage) + text{* This gives a simple derivation of limit component bounds. *} lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
--- a/src/HOL/Probability/Borel.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Borel.thy Fri Sep 03 13:54:39 2010 +0200 @@ -6,6 +6,10 @@ imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis begin +lemma LIMSEQ_max: + "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0" + by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) + section "Generic Borel spaces" definition "borel_space = sigma (UNIV::'a::topological_space set) open" @@ -81,7 +85,7 @@ "(\<lambda>x. c) \<in> borel_measurable M" by (auto intro!: measurable_const) -lemma (in sigma_algebra) borel_measurable_indicator: +lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]: assumes A: "A \<in> sets M" shows "indicator A \<in> borel_measurable M" unfolding indicator_def_raw using A @@ -105,6 +109,53 @@ qed (auto simp add: vimage_UN) qed +lemma (in sigma_algebra) borel_measurable_restricted: + fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M" + shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow> + (\<lambda>x. f x * indicator A x) \<in> borel_measurable M" + (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M") +proof - + interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) + have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R" + by (auto intro!: measurable_cong) + show ?thesis unfolding * + unfolding in_borel_measurable_borel_space + proof (simp, safe) + fix S :: "'x set" assume "S \<in> sets borel_space" + "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M" + then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto + then have f: "?f -` S \<inter> A \<in> sets M" + using `A \<in> sets M` sets_into_space by fastsimp + show "?f -` S \<inter> space M \<in> sets M" + proof cases + assume "0 \<in> S" + then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)" + using `A \<in> sets M` sets_into_space by auto + then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff) + next + assume "0 \<notin> S" + then have "?f -` S \<inter> space M = ?f -` S \<inter> A" + using `A \<in> sets M` sets_into_space + by (auto simp: indicator_def split: split_if_asm) + then show ?thesis using f by auto + qed + next + fix S :: "'x set" assume "S \<in> sets borel_space" + "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M" + then have f: "?f -` S \<inter> space M \<in> sets M" by auto + then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M" + using `A \<in> sets M` sets_into_space + apply (simp add: image_iff) + apply (rule bexI[OF _ f]) + by auto + qed +qed + +lemma (in sigma_algebra) borel_measurable_subalgebra: + assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)" + shows "f \<in> borel_measurable M" + using assms unfolding measurable_def by auto + section "Borel spaces on euclidean spaces" lemma lessThan_borel[simp, intro]: @@ -658,6 +709,30 @@ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)" using borel_measurable_iff_halfspace_greater[where 'c=real] by simp +lemma borel_measureable_euclidean_component: + "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel_space" + unfolding borel_space_def[where 'a=real] +proof (rule borel_space.measurable_sigma) + fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def . + from open_vimage_euclidean_component[OF this] + show "(\<lambda>x. x $$ i) -` S \<inter> space borel_space \<in> sets borel_space" + by (auto intro: borel_space_open) +qed auto + +lemma (in sigma_algebra) borel_measureable_euclidean_space: + fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space" + shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)" +proof safe + fix i assume "f \<in> borel_measurable M" + then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M" + using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def] + by (auto intro: borel_measureable_euclidean_component) +next + assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M" + then show "f \<in> borel_measurable M" + unfolding borel_measurable_iff_halfspace_le by auto +qed + subsection "Borel measurable operators" lemma (in sigma_algebra) affine_borel_measurable_vector: @@ -1270,4 +1345,46 @@ using assms by auto qed +lemma (in sigma_algebra) borel_measurable_psuminf: + assumes "\<And>i. f i \<in> borel_measurable M" + shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M" + using assms unfolding psuminf_def + by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) + +section "LIMSEQ is borel measurable" + +lemma (in sigma_algebra) borel_measurable_LIMSEQ: + fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" + assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" + and u: "\<And>i. u i \<in> borel_measurable M" + shows "u' \<in> borel_measurable M" +proof - + let "?pu x i" = "max (u i x) 0" + let "?nu x i" = "max (- u i x) 0" + + { fix x assume x: "x \<in> space M" + have "(?pu x) ----> max (u' x) 0" + "(?nu x) ----> max (- u' x) 0" + using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) + from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] + have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" + "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" + by (simp_all add: Real_max'[symmetric]) } + note eq = this + + have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x" + by auto + + have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M" + "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M" + using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real) + with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space] + have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M" + "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M" + unfolding SUPR_fun_expand INFI_fun_expand by auto + note this[THEN borel_measurable_real] + from borel_measurable_diff[OF this] + show ?thesis unfolding * . +qed + end
--- a/src/HOL/Probability/Caratheodory.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Fri Sep 03 13:54:39 2010 +0200 @@ -445,21 +445,6 @@ by intro_locales (auto simp add: sigma_algebra_def) qed - -lemma (in algebra) inf_measure_nonempty: - assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b" - shows "f b \<in> measure_set M f a" -proof - - have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}" - by (rule psuminf_finite) (simp add: f[unfolded positive_def]) - also have "... = f b" - by simp - finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" . - thus ?thesis using a b - by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"] - simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) -qed - lemma (in algebra) additive_increasing: assumes posf: "positive f" and addf: "additive M f" shows "increasing M f" @@ -494,6 +479,20 @@ by (auto simp add: Un binaryset_psuminf positive_def) qed +lemma inf_measure_nonempty: + assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M" + shows "f b \<in> measure_set M f a" +proof - + have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}" + by (rule psuminf_finite) (simp add: f[unfolded positive_def]) + also have "... = f b" + by simp + finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" . + thus ?thesis using assms + by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"] + simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) +qed + lemma (in algebra) inf_measure_agrees: assumes posf: "positive f" and ca: "countably_additive M f" and s: "s \<in> sets M" @@ -535,11 +534,11 @@ qed lemma (in algebra) inf_measure_empty: - assumes posf: "positive f" + assumes posf: "positive f" "{} \<in> sets M" shows "Inf (measure_set M f {}) = 0" proof (rule antisym) show "Inf (measure_set M f {}) \<le> 0" - by (metis complete_lattice_class.Inf_lower empty_sets inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) + by (metis complete_lattice_class.Inf_lower `{} \<in> sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) qed simp lemma (in algebra) inf_measure_positive: @@ -597,7 +596,7 @@ next case True have "measure_set M f s \<noteq> {}" - by (metis emptyE ss inf_measure_nonempty [of f, OF posf top]) + by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets]) then obtain l where "l \<in> measure_set M f s" by auto moreover from True have "l \<le> Inf (measure_set M f s) + e" by simp ultimately show ?thesis
--- a/src/HOL/Probability/Information.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Information.thy Fri Sep 03 13:54:39 2010 +0200 @@ -2,11 +2,53 @@ imports Probability_Space Product_Measure Convex Radon_Nikodym begin +lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y" + by (subst log_le_cancel_iff) auto + +lemma log_less: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x < y \<Longrightarrow> log a x < log a y" + by (subst log_less_cancel_iff) auto + +lemma setsum_cartesian_product': + "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)" + unfolding setsum_cartesian_product by simp + lemma real_of_pinfreal_inverse[simp]: fixes X :: pinfreal shows "real (inverse X) = 1 / real X" by (cases X) (auto simp: inverse_eq_divide) +lemma (in finite_prob_space) finite_product_prob_space_of_images: + "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr> + (joint_distribution X Y)" + (is "finite_prob_space ?S _") +proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) + have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto + thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1" + by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) +qed + +lemma (in finite_prob_space) finite_measure_space_prod: + assumes X: "finite_measure_space MX (distribution X)" + assumes Y: "finite_measure_space MY (distribution Y)" + shows "finite_measure_space (prod_measure_space MX MY) (joint_distribution X Y)" + (is "finite_measure_space ?M ?D") +proof (intro finite_measure_spaceI) + interpret X: finite_measure_space MX "distribution X" by fact + interpret Y: finite_measure_space MY "distribution Y" by fact + note finite_measure_space.finite_prod_measure_space[OF X Y, simp] + show "finite (space ?M)" using X.finite_space Y.finite_space by auto + show "joint_distribution X Y {} = 0" by simp + show "sets ?M = Pow (space ?M)" by simp + { fix x show "?D (space ?M) \<noteq> \<omega>" by (rule distribution_finite) } + { fix A B assume "A \<subseteq> space ?M" "B \<subseteq> space ?M" "A \<inter> B = {}" + have *: "(\<lambda>t. (X t, Y t)) -` (A \<union> B) \<inter> space M = + (\<lambda>t. (X t, Y t)) -` A \<inter> space M \<union> (\<lambda>t. (X t, Y t)) -` B \<inter> space M" + by auto + show "?D (A \<union> B) = ?D A + ?D B" unfolding distribution_def * + apply (rule measure_additive[symmetric]) + using `A \<inter> B = {}` by (auto simp: sets_eq_Pow) } +qed + section "Convex theory" lemma log_setsum: @@ -105,51 +147,19 @@ finally show ?thesis . qed -lemma (in finite_prob_space) sum_over_space_distrib: - "(\<Sum>x\<in>X`space M. distribution X {x}) = 1" - unfolding distribution_def measure_space_1[symmetric] using finite_space - by (subst measure_finitely_additive'') - (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>]) - -lemma (in finite_prob_space) sum_over_space_real_distribution: - "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1" - unfolding distribution_def prob_space[symmetric] using finite_space - by (subst real_finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob]) +lemma split_pairs: + shows + "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and + "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto section "Information theory" -definition - "KL_divergence b M \<mu> \<nu> = - measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))" - locale finite_information_space = finite_prob_space + fixes b :: real assumes b_gt_1: "1 < b" -lemma (in finite_prob_space) distribution_mono: - assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" - shows "distribution X x \<le> distribution Y y" - unfolding distribution_def - using assms by (auto simp: sets_eq_Pow intro!: measure_mono) - -lemma (in prob_space) distribution_remove_const: - shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}" - and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}" - and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" - and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" - and "distribution (\<lambda>x. ()) {()} = 1" - unfolding measure_space_1[symmetric] - by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def) - context finite_information_space begin -lemma distribution_mono_gt_0: - assumes gt_0: "0 < distribution X x" - assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" - shows "0 < distribution Y y" - by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) - lemma assumes "0 \<le> A" and pos: "0 < A \<Longrightarrow> 0 < B" "0 < A \<Longrightarrow> 0 < C" shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult") @@ -165,41 +175,6 @@ thus ?mult and ?div by auto qed -lemma split_pairs: - shows - "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and - "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto - -lemma (in finite_information_space) distribution_finite: - "distribution X A \<noteq> \<omega>" - using measure_finite[of "X -` A \<inter> space M"] - unfolding distribution_def sets_eq_Pow by auto - -lemma (in finite_information_space) real_distribution_gt_0[simp]: - "0 < real (distribution Y y) \<longleftrightarrow> 0 < distribution Y y" - using assms by (auto intro!: real_pinfreal_pos distribution_finite) - -lemma real_distribution_mult_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * distribution X x)" - unfolding real_of_pinfreal_mult[symmetric] - using assms by (auto intro!: mult_pos_pos) - -lemma real_distribution_divide_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y / distribution X x)" - unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) - -lemma real_distribution_mult_inverse_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * inverse (distribution X x))" - unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) - ML {* (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"} @@ -252,31 +227,14 @@ end -lemma (in finite_measure_space) absolutely_continuousI: - assumes "finite_measure_space M \<nu>" - assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0" - shows "absolutely_continuous \<nu>" -proof (unfold absolutely_continuous_def sets_eq_Pow, safe) - fix N assume "\<mu> N = 0" "N \<subseteq> space M" - - interpret v: finite_measure_space M \<nu> by fact +subsection "Kullback$-$Leibler divergence" - have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp - also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})" - proof (rule v.measure_finitely_additive''[symmetric]) - show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset) - show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto - fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto - qed - also have "\<dots> = 0" - proof (safe intro!: setsum_0') - fix x assume "x \<in> N" - hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono) - hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp - thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto - qed - finally show "\<nu> N = 0" . -qed +text {* The Kullback$-$Leibler divergence is also known as relative entropy or +Kullback$-$Leibler distance. *} + +definition + "KL_divergence b M \<mu> \<nu> = + measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))" lemma (in finite_measure_space) KL_divergence_eq_finite: assumes v: "finite_measure_space M \<nu>" @@ -285,19 +243,13 @@ proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v]) interpret v: finite_measure_space M \<nu> by fact have ms: "measure_space M \<nu>" by fact - have ac: "absolutely_continuous \<nu>" using ac by (auto intro!: absolutely_continuousI[OF v]) - show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum" using RN_deriv_finite_measure[OF ms ac] by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric]) qed -lemma (in finite_prob_space) finite_sum_over_space_eq_1: - "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1" - using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow) - lemma (in finite_prob_space) KL_divergence_positive_finite: assumes v: "finite_prob_space M \<nu>" assumes ac: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0" @@ -322,13 +274,15 @@ fix x assume x: "x \<in> space M" { assume "0 < real (\<nu> {x})" hence "\<mu> {x} \<noteq> 0" using ac[OF x] by auto - thus "0 < prob {x}" using measure_finite[of "{x}"] sets_eq_Pow x + thus "0 < prob {x}" using finite_measure[of "{x}"] sets_eq_Pow x by (cases "\<mu> {x}") simp_all } qed auto qed thus "0 \<le> KL_divergence b M \<nu> \<mu>" using finite_sum_over_space_eq_1 by simp qed +subsection {* Mutual Information *} + definition (in prob_space) "mutual_information b S T X Y = KL_divergence b (prod_measure_space S T) @@ -341,24 +295,48 @@ \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y" -lemma prod_measure_times_finite: - assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N" - shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}" -proof (cases a) - case (Pair b c) - hence a_eq: "{a} = {b} \<times> {c}" by simp - - interpret M: finite_measure_space M \<mu> by fact - interpret N: finite_measure_space N \<nu> by fact - - from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair - show ?thesis unfolding a_eq by simp +lemma (in finite_information_space) mutual_information_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MY (distribution Y)" + shows "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)}) / + (real (distribution X {x}) * real (distribution Y {y}))))" +proof - + let ?P = "prod_measure_space MX MY" + let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)" + let ?\<nu> = "joint_distribution X Y" + interpret X: finite_measure_space MX "distribution X" by fact + moreover interpret Y: finite_measure_space MY "distribution Y" by fact + have fms: "finite_measure_space MX (distribution X)" + "finite_measure_space MY (distribution Y)" by fact+ + have fms_P: "finite_measure_space ?P ?\<mu>" + by (rule X.finite_measure_space_finite_prod_measure) fact + then interpret P: finite_measure_space ?P ?\<mu> . + have fms_P': "finite_measure_space ?P ?\<nu>" + using finite_product_measure_space[of "space MX" "space MY"] + X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space] + X.sets_eq_Pow Y.sets_eq_Pow + by (simp add: prod_measure_space_def sigma_def) + then interpret P': finite_measure_space ?P ?\<nu> . + { fix x assume "x \<in> space ?P" + hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow + by (auto simp: prod_measure_space_def) + assume "?\<mu> {x} = 0" + with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX + have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0" + by (simp add: prod_measure_space_def) + hence "joint_distribution X Y {x} = 0" + by (cases x) (auto simp: distribution_order) } + note measure_0 = this + show ?thesis + unfolding Let_def mutual_information_def + using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def + by (subst P.KL_divergence_eq_finite) + (auto simp add: prod_measure_space_def prod_measure_times_finite + finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric]) qed -lemma setsum_cartesian_product': - "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)" - unfolding setsum_cartesian_product by simp - lemma (in finite_information_space) assumes MX: "finite_prob_space MX (distribution X)" assumes MY: "finite_prob_space MY (distribution Y)" @@ -436,9 +414,26 @@ (real (distribution X {x}) * real (distribution Y {y}))))" by (subst mutual_information_eq_generic) (simp_all add: finite_prob_space_of_images) +lemma (in finite_information_space) mutual_information_cong: + assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" + assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" + shows "\<I>(X ; Y) = \<I>(X' ; Y')" +proof - + have "X ` space M = X' ` space M" using X by (auto intro!: image_eqI) + moreover have "Y ` space M = Y' ` space M" using Y by (auto intro!: image_eqI) + ultimately show ?thesis + unfolding mutual_information_eq + using + assms[THEN distribution_cong] + joint_distribution_cong[OF assms] + by (auto intro!: setsum_cong) +qed + lemma (in finite_information_space) mutual_information_positive: "0 \<le> \<I>(X;Y)" by (subst mutual_information_positive_generic) (simp_all add: finite_prob_space_of_images) +subsection {* Entropy *} + definition (in prob_space) "entropy b s X = mutual_information b s s X X" @@ -446,32 +441,146 @@ finite_entropy ("\<H>'(_')") where "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X" -lemma (in finite_information_space) joint_distribution_remove[simp]: - "joint_distribution X X {(x, x)} = distribution X {x}" - unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"]) +lemma (in finite_information_space) entropy_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + shows "entropy b MX X = -(\<Sum> x \<in> space MX. real (distribution X {x}) * log b (real (distribution X {x})))" +proof - + let "?X x" = "real (distribution X {x})" + let "?XX x y" = "real (joint_distribution X X {(x, y)})" + interpret MX: finite_measure_space MX "distribution X" by fact + { fix x y + have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto + then have "?XX x y * log b (?XX x y / (?X x * ?X y)) = + (if x = y then - ?X y * log b (?X y) else 0)" + unfolding distribution_def by (auto simp: mult_log_divide) } + note remove_XX = this + show ?thesis + unfolding entropy_def mutual_information_generic_eq[OF MX MX] + unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX + by (auto simp: setsum_cases MX.finite_space) +qed lemma (in finite_information_space) entropy_eq: "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))" -proof - - { fix f - { fix x y - have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto - hence "real (distribution (\<lambda>x. (X x, X x)) {(x,y)}) * f x y = - (if x = y then real (distribution X {x}) * f x y else 0)" - unfolding distribution_def by auto } - hence "(\<Sum>(x, y) \<in> X ` space M \<times> X ` space M. real (joint_distribution X X {(x, y)}) * f x y) = - (\<Sum>x \<in> X ` space M. real (distribution X {x}) * f x x)" - unfolding setsum_cartesian_product' by (simp add: setsum_cases finite_space) } - note remove_cartesian_product = this - - show ?thesis - unfolding entropy_def mutual_information_eq setsum_negf[symmetric] remove_cartesian_product - by (auto intro!: setsum_cong) -qed + by (simp add: finite_measure_space entropy_generic_eq) lemma (in finite_information_space) entropy_positive: "0 \<le> \<H>(X)" unfolding entropy_def using mutual_information_positive . +lemma (in finite_information_space) entropy_certainty_eq_0: + assumes "x \<in> X ` space M" and "distribution X {x} = 1" + shows "\<H>(X) = 0" +proof - + interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X" + by (rule finite_prob_space_of_images) + + have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" + using X.measure_compl[of "{x}"] assms by auto + also have "\<dots> = 0" using X.prob_space assms by auto + finally have X0: "distribution X (X ` space M - {x}) = 0" by auto + + { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M" + hence "{y} \<subseteq> X ` space M - {x}" by auto + from X.measure_mono[OF this] X0 asm + have "distribution X {y} = 0" by auto } + + hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)" + using assms by auto + + have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp + + show ?thesis unfolding entropy_eq by (auto simp: y fi) +qed + +lemma (in finite_information_space) entropy_le_card_not_0: + "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))" +proof - + let "?d x" = "distribution X {x}" + let "?p x" = "real (?d x)" + have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))" + by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric]) + also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))" + apply (rule log_setsum') + using not_empty b_gt_1 finite_space sum_over_space_real_distribution + by auto + also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)" + apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"]) + using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0) + finally show ?thesis + using finite_space by (auto simp: setsum_cases real_eq_of_nat) +qed + +lemma (in finite_information_space) entropy_uniform_max: + assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}" + shows "\<H>(X) = log b (real (card (X ` space M)))" +proof - + note uniform = + finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified] + + have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff + using finite_space not_empty by auto + + { fix x assume "x \<in> X ` space M" + hence "real (distribution X {x}) = 1 / real (card (X ` space M))" + proof (rule uniform) + fix x y assume "x \<in> X`space M" "y \<in> X`space M" + from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp + qed } + thus ?thesis + using not_empty finite_space b_gt_1 card_gt0 + by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide) +qed + +lemma (in finite_information_space) entropy_le_card: + "\<H>(X) \<le> log b (real (card (X ` space M)))" +proof cases + assume "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} = {}" + then have "\<And>x. x\<in>X`space M \<Longrightarrow> distribution X {x} = 0" by auto + moreover + have "0 < card (X`space M)" + using finite_space not_empty unfolding card_gt_0_iff by auto + then have "log b 1 \<le> log b (real (card (X`space M)))" + using b_gt_1 by (intro log_le) auto + ultimately show ?thesis unfolding entropy_eq by simp +next + assume False: "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} \<noteq> {}" + have "card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}) \<le> card (X ` space M)" + (is "?A \<le> ?B") using finite_space not_empty by (auto intro!: card_mono) + note entropy_le_card_not_0 + also have "log b (real ?A) \<le> log b (real ?B)" + using b_gt_1 False finite_space not_empty `?A \<le> ?B` + by (auto intro!: log_le simp: card_gt_0_iff) + finally show ?thesis . +qed + +lemma (in finite_information_space) entropy_commute: + "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))" +proof - + have *: "(\<lambda>x. (Y x, X x))`space M = (\<lambda>(a,b). (b,a))`(\<lambda>x. (X x, Y x))`space M" + by auto + have inj: "\<And>X. inj_on (\<lambda>(a,b). (b,a)) X" + by (auto intro!: inj_onI) + show ?thesis + unfolding entropy_eq unfolding * setsum_reindex[OF inj] + by (simp add: joint_distribution_commute[of Y X] split_beta) +qed + +lemma (in finite_information_space) entropy_eq_cartesian_sum: + "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)})))" +proof - + { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M" + then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto + then have "joint_distribution X Y {x} = 0" + unfolding distribution_def by auto } + then show ?thesis using finite_space + unfolding entropy_eq neg_equal_iff_equal setsum_cartesian_product + by (auto intro!: setsum_mono_zero_cong_left) +qed + +subsection {* Conditional Mutual Information *} + definition (in prob_space) "conditional_mutual_information b M1 M2 M3 X Y Z \<equiv> mutual_information b M1 (prod_measure_space M2 M3) X (\<lambda>x. (Y x, Z x)) - @@ -485,87 +594,32 @@ \<lparr> space = Z`space M, sets = Pow (Z`space M) \<rparr> X Y Z" -lemma (in finite_information_space) setsum_distribution_gen: - assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M" - and "inj_on f (X`space M)" - shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}" - unfolding distribution_def assms - using finite_space assms - by (subst measure_finitely_additive'') - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_information_space) setsum_distribution: - "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" - "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" - "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" - "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" - "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" - by (auto intro!: inj_onI setsum_distribution_gen) - -lemma (in finite_information_space) setsum_real_distribution_gen: - assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M" - and "inj_on f (X`space M)" - shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})" - unfolding distribution_def assms - using finite_space assms - by (subst real_finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_information_space) setsum_real_distribution: - "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})" - "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})" - "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})" - "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})" - "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})" - by (auto intro!: inj_onI setsum_real_distribution_gen) +lemma (in finite_information_space) conditional_mutual_information_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MY (distribution Y)" + assumes MZ: "finite_measure_space MZ (distribution Z)" + shows "conditional_mutual_information b MX MY MZ X Y Z = + (\<Sum>(x, y, z)\<in>space MX \<times> space MY \<times> space MZ. + real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) * + log b (real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) / + (real (distribution X {x}) * real (joint_distribution Y Z {(y, z)})))) - + (\<Sum>(x, y)\<in>space MX \<times> space MZ. + real (joint_distribution X Z {(x, y)}) * + log b (real (joint_distribution X Z {(x, y)}) / (real (distribution X {x}) * real (distribution Z {y}))))" + using assms finite_measure_space_prod[OF MY MZ] + unfolding conditional_mutual_information_def + by (subst (1 2) mutual_information_generic_eq) + (simp_all add: setsum_cartesian_product' finite_measure_space.finite_prod_measure_space) -lemma (in finite_information_space) conditional_mutual_information_eq_sum: - "\<I>(X ; Y | Z) = - (\<Sum>(x, y, z)\<in>X ` space M \<times> (\<lambda>x. (Y x, Z x)) ` space M. - real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) * - log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)})/ - real (distribution (\<lambda>x. (Y x, Z x)) {(y, z)}))) - - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M. - real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) * log b (real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) / real (distribution Z {z})))" - (is "_ = ?rhs") -proof - - have setsum_product: - "\<And>f x. (\<Sum>v\<in>(\<lambda>x. (Y x, Z x)) ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v) - = (\<Sum>v\<in>Y ` space M \<times> Z ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v)" - proof (safe intro!: setsum_mono_zero_cong_left imageI) - fix x y z f - assume *: "(Y y, Z z) \<notin> (\<lambda>x. (Y x, Z x)) ` space M" and "y \<in> space M" "z \<in> space M" - hence "(\<lambda>x. (X x, Y x, Z x)) -` {(x, Y y, Z z)} \<inter> space M = {}" - proof safe - fix x' assume x': "x' \<in> space M" and eq: "Y x' = Y y" "Z x' = Z z" - have "(Y y, Z z) \<in> (\<lambda>x. (Y x, Z x)) ` space M" using eq[symmetric] x' by auto - thus "x' \<in> {}" using * by auto - qed - thus "real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, Y y, Z z)}) * f (Y y) (Z z) = 0" - unfolding distribution_def by simp - qed (simp add: finite_space) - - thus ?thesis - unfolding conditional_mutual_information_def Let_def mutual_information_eq - by (subst mutual_information_eq_generic) - (auto simp: prod_measure_space_def sigma_prod_sets_finite finite_space sigma_def - finite_prob_space_of_images finite_product_prob_space_of_images - setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf - setsum_left_distrib[symmetric] setsum_real_distribution - cong: setsum_cong) -qed lemma (in finite_information_space) conditional_mutual_information_eq: "\<I>(X ; Y | Z) = (\<Sum>(x, y, z) \<in> X ` space M \<times> Y ` space M \<times> Z ` space M. real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) * log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) / (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" - unfolding conditional_mutual_information_def Let_def mutual_information_eq - by (subst mutual_information_eq_generic) + by (subst conditional_mutual_information_generic_eq) (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space - finite_prob_space_of_images finite_product_prob_space_of_images sigma_def + finite_measure_space finite_product_prob_space_of_images sigma_def setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf setsum_left_distrib[symmetric] setsum_real_distribution setsum_commute[where A="Y`space M"] real_of_pinfreal_mult[symmetric] @@ -581,22 +635,6 @@ by (simp add: setsum_cartesian_product' distribution_remove_const) qed -lemma (in finite_prob_space) distribution_finite: - "distribution X A \<noteq> \<omega>" - by (auto simp: sets_eq_Pow distribution_def intro!: measure_finite) - -lemma (in finite_prob_space) real_distribution_order: - shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})" - and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})" - and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})" - and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})" - and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" - and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" - using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] - using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] - using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] - by auto - lemma (in finite_information_space) conditional_mutual_information_positive: "0 \<le> \<I>(X ; Y | Z)" proof - @@ -640,6 +678,8 @@ by (simp add: real_of_pinfreal_mult[symmetric]) qed +subsection {* Conditional Entropy *} + definition (in prob_space) "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y" @@ -652,19 +692,69 @@ lemma (in finite_information_space) conditional_entropy_positive: "0 \<le> \<H>(X | Y)" unfolding conditional_entropy_def using conditional_mutual_information_positive . +lemma (in finite_information_space) conditional_entropy_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MZ (distribution Z)" + shows "conditional_entropy b MX MZ X Z = + - (\<Sum>(x, z)\<in>space MX \<times> space MZ. + real (joint_distribution X Z {(x, z)}) * + log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" + unfolding conditional_entropy_def using assms + apply (simp add: conditional_mutual_information_generic_eq + setsum_cartesian_product' setsum_commute[of _ "space MZ"] + setsum_negf[symmetric] setsum_subtractf[symmetric]) +proof (safe intro!: setsum_cong, simp) + fix z x assume "z \<in> space MZ" "x \<in> space MX" + let "?XXZ x'" = "real (joint_distribution X (\<lambda>x. (X x, Z x)) {(x, x', z)})" + let "?XZ x'" = "real (joint_distribution X Z {(x', z)})" + let "?X" = "real (distribution X {x})" + interpret MX: finite_measure_space MX "distribution X" by fact + have *: "\<And>A. A = {} \<Longrightarrow> prob A = 0" by simp + have XXZ: "\<And>x'. ?XXZ x' = (if x' = x then ?XZ x else 0)" + by (auto simp: distribution_def intro!: arg_cong[where f=prob] *) + have "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) = + (\<Sum>x'\<in>{x}. ?XZ x' * log b (?XZ x') - (?XZ x' * log b ?X + ?XZ x' * log b (?XZ x')))" + using `x \<in> space MX` MX.finite_space + by (safe intro!: setsum_mono_zero_cong_right) + (auto split: split_if_asm simp: XXZ) + then show "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) + + ?XZ x * log b ?X = 0" by simp +qed + lemma (in finite_information_space) conditional_entropy_eq: "\<H>(X | Z) = - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M. real (joint_distribution X Z {(x, z)}) * log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" + by (simp add: finite_measure_space conditional_entropy_generic_eq) + +lemma (in finite_information_space) conditional_entropy_eq_ce_with_hypothesis: + "\<H>(X | Y) = + -(\<Sum>y\<in>Y`space M. real (distribution Y {y}) * + (\<Sum>x\<in>X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) * + log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))" + unfolding conditional_entropy_eq neg_equal_iff_equal + apply (simp add: setsum_commute[of _ "Y`space M"] setsum_cartesian_product' setsum_divide_distrib[symmetric]) + apply (safe intro!: setsum_cong) + using real_distribution_order'[of Y _ X _] + by (auto simp add: setsum_subtractf[of _ _ "X`space M"]) + +lemma (in finite_information_space) conditional_entropy_eq_cartesian_sum: + "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))" proof - - have *: "\<And>x y z. (\<lambda>x. (X x, X x, Z x)) -` {(x, y, z)} = (if x = y then (\<lambda>x. (X x, Z x)) -` {(x, z)} else {})" by auto - show ?thesis - unfolding conditional_mutual_information_eq_sum - conditional_entropy_def distribution_def * - by (auto intro!: setsum_0') + { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M" + then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto + then have "joint_distribution X Y {x} = 0" + unfolding distribution_def by auto } + then show ?thesis using finite_space + unfolding conditional_entropy_eq neg_equal_iff_equal setsum_cartesian_product + by (auto intro!: setsum_mono_zero_cong_left) qed +subsection {* Equalities *} + lemma (in finite_information_space) mutual_information_eq_entropy_conditional_entropy: "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" unfolding mutual_information_eq entropy_eq conditional_entropy_eq @@ -680,109 +770,15 @@ show ?thesis by auto qed -(* -------------Entropy of a RV with a certain event is zero---------------- *) - -lemma (in finite_information_space) finite_entropy_certainty_eq_0: - assumes "x \<in> X ` space M" and "distribution X {x} = 1" - shows "\<H>(X) = 0" -proof - - interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X" - by (rule finite_prob_space_of_images) - - have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" - using X.measure_compl[of "{x}"] assms by auto - also have "\<dots> = 0" using X.prob_space assms by auto - finally have X0: "distribution X (X ` space M - {x}) = 0" by auto - - { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M" - hence "{y} \<subseteq> X ` space M - {x}" by auto - from X.measure_mono[OF this] X0 asm - have "distribution X {y} = 0" by auto } - - hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)" - using assms by auto - - have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp - - show ?thesis unfolding entropy_eq by (auto simp: y fi) -qed -(* --------------- upper bound on entropy for a rv ------------------------- *) - -lemma (in finite_prob_space) distribution_1: - "distribution X A \<le> 1" - unfolding distribution_def measure_space_1[symmetric] - by (auto intro!: measure_mono simp: sets_eq_Pow) - -lemma (in finite_prob_space) real_distribution_1: - "real (distribution X A) \<le> 1" - unfolding real_pinfreal_1[symmetric] - by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp +lemma (in finite_information_space) entropy_chain_rule: + "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)" + unfolding entropy_eq[of X] entropy_eq_cartesian_sum conditional_entropy_eq_cartesian_sum + unfolding setsum_commute[of _ "X`space M"] setsum_negf[symmetric] setsum_addf[symmetric] + by (rule setsum_cong) + (simp_all add: setsum_negf setsum_addf setsum_subtractf setsum_real_distribution + setsum_left_distrib[symmetric] joint_distribution_commute[of X Y]) -lemma (in finite_information_space) finite_entropy_le_card: - "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))" -proof - - let "?d x" = "distribution X {x}" - let "?p x" = "real (?d x)" - have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))" - by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric]) - also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))" - apply (rule log_setsum') - using not_empty b_gt_1 finite_space sum_over_space_real_distribution - by auto - also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)" - apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"]) - using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0) - finally show ?thesis - using finite_space by (auto simp: setsum_cases real_eq_of_nat) -qed - -(* --------------- entropy is maximal for a uniform rv --------------------- *) - -lemma (in finite_prob_space) uniform_prob: - assumes "x \<in> space M" - assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}" - shows "prob {x} = 1 / real (card (space M))" -proof - - have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}" - using assms(2)[OF _ `x \<in> space M`] by blast - have "1 = prob (space M)" - using prob_space by auto - also have "\<dots> = (\<Sum> x \<in> space M. prob {x})" - using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified] - sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] - finite_space unfolding disjoint_family_on_def prob_space[symmetric] - by (auto simp add:setsum_restrict_set) - also have "\<dots> = (\<Sum> y \<in> space M. prob {x})" - using prob_x by auto - also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp - finally have one: "1 = real (card (space M)) * prob {x}" - using real_eq_of_nat by auto - hence two: "real (card (space M)) \<noteq> 0" by fastsimp - from one have three: "prob {x} \<noteq> 0" by fastsimp - thus ?thesis using one two three divide_cancel_right - by (auto simp:field_simps) -qed - -lemma (in finite_information_space) finite_entropy_uniform_max: - assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}" - shows "\<H>(X) = log b (real (card (X ` space M)))" -proof - - note uniform = - finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified] - - have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff - using finite_space not_empty by auto - - { fix x assume "x \<in> X ` space M" - hence "real (distribution X {x}) = 1 / real (card (X ` space M))" - proof (rule uniform) - fix x y assume "x \<in> X`space M" "y \<in> X`space M" - from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp - qed } - thus ?thesis - using not_empty finite_space b_gt_1 card_gt0 - by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide) -qed +section {* Partitioning *} definition "subvimage A f g \<longleftrightarrow> (\<forall>x \<in> A. f -` {f x} \<inter> A \<subseteq> g -` {g x} \<inter> A)" @@ -934,38 +930,6 @@ "\<H>(f \<circ> X) \<le> \<H>(X)" by (subst (2) entropy_partition[of _ "f \<circ> X"]) (auto intro: conditional_entropy_positive) -lemma (in prob_space) distribution_cong: - assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x" - shows "distribution X = distribution Y" - unfolding distribution_def expand_fun_eq - using assms by (auto intro!: arg_cong[where f="\<mu>"]) - -lemma (in prob_space) joint_distribution_cong: - assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" - assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" - shows "joint_distribution X Y = joint_distribution X' Y'" - unfolding distribution_def expand_fun_eq - using assms by (auto intro!: arg_cong[where f="\<mu>"]) - -lemma image_cong: - "\<lbrakk> \<And>x. x \<in> S \<Longrightarrow> X x = X' x \<rbrakk> \<Longrightarrow> X ` S = X' ` S" - by (auto intro!: image_eqI) - -lemma (in finite_information_space) mutual_information_cong: - assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" - assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" - shows "\<I>(X ; Y) = \<I>(X' ; Y')" -proof - - have "X ` space M = X' ` space M" using X by (rule image_cong) - moreover have "Y ` space M = Y' ` space M" using Y by (rule image_cong) - ultimately show ?thesis - unfolding mutual_information_eq - using - assms[THEN distribution_cong] - joint_distribution_cong[OF assms] - by (auto intro!: setsum_cong) -qed - corollary (in finite_information_space) entropy_of_inj: assumes "inj_on f (X`space M)" shows "\<H>(f \<circ> X) = \<H>(X)"
--- a/src/HOL/Probability/Lebesgue_Integration.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Fri Sep 03 13:54:39 2010 +0200 @@ -209,19 +209,6 @@ by (auto intro!: **) qed -lemma setsum_indicator_disjoint_family: - fixes f :: "'d \<Rightarrow> 'e::semiring_1" - assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P" - shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j" -proof - - have "P \<inter> {i. x \<in> A i} = {j}" - using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def - by auto - thus ?thesis - unfolding indicator_def - by (simp add: if_distrib setsum_cases[OF `finite P`]) -qed - lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: fixes u :: "'a \<Rightarrow> pinfreal" assumes u: "u \<in> borel_measurable M" @@ -426,6 +413,62 @@ with x show thesis by (auto intro!: that[of f]) qed +lemma (in sigma_algebra) simple_function_eq_borel_measurable: + fixes f :: "'a \<Rightarrow> pinfreal" + shows "simple_function f \<longleftrightarrow> + finite (f`space M) \<and> f \<in> borel_measurable M" + using simple_function_borel_measurable[of f] + borel_measurable_simple_function[of f] + by (fastsimp simp: simple_function_def) + +lemma (in measure_space) simple_function_restricted: + fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M" + shows "sigma_algebra.simple_function (restricted_space A) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)" + (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f") +proof - + interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) + have "finite (f`A) \<longleftrightarrow> finite (?f`space M)" + proof cases + assume "A = space M" + then have "f`A = ?f`space M" by (fastsimp simp: image_iff) + then show ?thesis by simp + next + assume "A \<noteq> space M" + then obtain x where x: "x \<in> space M" "x \<notin> A" + using sets_into_space `A \<in> sets M` by auto + have *: "?f`space M = f`A \<union> {0}" + proof (auto simp add: image_iff) + show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0" + using x by (auto intro!: bexI[of _ x]) + next + fix x assume "x \<in> A" + then show "\<exists>y\<in>space M. f x = f y * indicator A y" + using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x]) + next + fix x + assume "indicator A x \<noteq> (0::pinfreal)" + then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm) + moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y" + ultimately show "f x = 0" by auto + qed + then show ?thesis by auto + qed + then show ?thesis + unfolding simple_function_eq_borel_measurable + R.simple_function_eq_borel_measurable + unfolding borel_measurable_restricted[OF `A \<in> sets M`] + by auto +qed + +lemma (in sigma_algebra) simple_function_subalgebra: + assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f" + and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)" + shows "simple_function f" + using assms + unfolding simple_function_def + unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)] + by auto + section "Simple integral" definition (in measure_space) @@ -668,6 +711,41 @@ qed qed +lemma (in measure_space) simple_integral_restricted: + assumes "A \<in> sets M" + assumes sf: "simple_function (\<lambda>x. f x * indicator A x)" + shows "measure_space.simple_integral (restricted_space A) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)" + (is "_ = simple_integral ?f") + unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]] + unfolding simple_integral_def +proof (simp, safe intro!: setsum_mono_zero_cong_left) + from sf show "finite (?f ` space M)" + unfolding simple_function_def by auto +next + fix x assume "x \<in> A" + then show "f x \<in> ?f ` space M" + using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x]) +next + fix x assume "x \<in> space M" "?f x \<notin> f`A" + then have "x \<notin> A" by (auto simp: image_iff) + then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp +next + fix x assume "x \<in> A" + then have "f x \<noteq> 0 \<Longrightarrow> + f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M" + using `A \<in> sets M` sets_into_space + by (auto simp: indicator_def split: split_if_asm) + then show "f x * \<mu> (f -` {f x} \<inter> A) = + f x * \<mu> (?f -` {f x} \<inter> space M)" + unfolding pinfreal_mult_cancel_left by auto +qed + +lemma (in measure_space) simple_integral_subalgebra[simp]: + assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>" + shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral" + unfolding simple_integral_def_raw + unfolding measure_space.simple_integral_def_raw[OF assms] by simp + section "Continuous posititve integration" definition (in measure_space) @@ -1077,6 +1155,43 @@ qed qed +lemma (in measure_space) positive_integral_translated_density: + assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" + shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g = + positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") +proof - + from measure_space_density[OF assms(1)] + interpret T: measure_space M ?T . + from borel_measurable_implies_simple_function_sequence[OF assms(2)] + obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast + note G_borel = borel_measurable_simple_function[OF this(1)] + from T.positive_integral_isoton[OF `G \<up> g` G_borel] + have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" . + { fix i + have [simp]: "finite (G i ` space M)" + using G(1) unfolding simple_function_def by auto + have "T.positive_integral (G i) = T.simple_integral (G i)" + using G T.positive_integral_eq_simple_integral by simp + also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))" + apply (simp add: T.simple_integral_def) + apply (subst positive_integral_cmult[symmetric]) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + apply (subst positive_integral_setsum[symmetric]) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + by (simp add: setsum_right_distrib field_simps) + also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)" + by (auto intro!: positive_integral_cong + simp: indicator_def if_distrib setsum_cases) + finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . } + with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp + from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)" + unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) + then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)" + using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times) + with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)" + unfolding isoton_def by simp +qed + lemma (in measure_space) positive_integral_null_set: assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets" shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0") @@ -1222,6 +1337,58 @@ finally show ?thesis by simp qed +lemma (in measure_space) positive_integral_restricted: + assumes "A \<in> sets M" + shows "measure_space.positive_integral (restricted_space A) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)" + (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f") +proof - + have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`]) + then interpret R: measure_space ?R \<mu> . + have saR: "sigma_algebra ?R" by fact + have *: "R.positive_integral f = R.positive_integral ?f" + by (auto intro!: R.positive_integral_cong) + show ?thesis + unfolding * R.positive_integral_def positive_integral_def + unfolding simple_function_restricted[OF `A \<in> sets M`] + apply (simp add: SUPR_def) + apply (rule arg_cong[where f=Sup]) + proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`]) + fix g assume "simple_function (\<lambda>x. g x * indicator A x)" + "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x" + then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and> + simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x" + apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) + by (auto simp: indicator_def le_fun_def) + next + fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)" + "\<forall>x\<in>space M. \<omega> \<noteq> g x" + then have *: "(\<lambda>x. g x * indicator A x) = g" + "\<And>x. g x * indicator A x = g x" + "\<And>x. g x \<le> f x" + by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm) + from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and> + simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)" + using `A \<in> sets M`[THEN sets_into_space] + apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) + by (fastsimp simp: le_fun_def *) + qed +qed + +lemma (in measure_space) positive_integral_subalgebra[simp]: + assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)" + and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)" + shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f" +proof - + note msN = measure_space_subalgebra[OF N_subalgebra] + then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> . + from N.borel_measurable_implies_simple_function_sequence[OF borel] + obtain fs where Nsf: "\<And>i. N.simple_function (fs i)" and "fs \<up> f" by blast + then have sf: "\<And>i. simple_function (fs i)" + using simple_function_subalgebra[OF _ N_subalgebra] by blast + from positive_integral_isoton_simple[OF `fs \<up> f` sf] N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf] + show ?thesis unfolding simple_integral_subalgebra[OF msN] isoton_def by simp +qed + section "Lebesgue Integral" definition (in measure_space) integrable where @@ -1629,44 +1796,6 @@ by (simp add: real_of_pinfreal_eq_0) qed -lemma LIMSEQ_max: - "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0" - by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) - -lemma (in sigma_algebra) borel_measurable_LIMSEQ: - fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real" - assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" - and u: "\<And>i. u i \<in> borel_measurable M" - shows "u' \<in> borel_measurable M" -proof - - let "?pu x i" = "max (u i x) 0" - let "?nu x i" = "max (- u i x) 0" - - { fix x assume x: "x \<in> space M" - have "(?pu x) ----> max (u' x) 0" - "(?nu x) ----> max (- u' x) 0" - using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) - from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] - have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" - "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" - by (simp_all add: Real_max'[symmetric]) } - note eq = this - - have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x" - by auto - - have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M" - "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M" - using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real) - with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space] - have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M" - "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M" - unfolding SUPR_fun_expand INFI_fun_expand by auto - note this[THEN borel_measurable_real] - from borel_measurable_diff[OF this] - show ?thesis unfolding * . -qed - lemma (in measure_space) integral_dominated_convergence: assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x" and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x" @@ -1926,41 +2055,11 @@ by (simp_all add: integral_cmul_indicator borel_measurable_vimage) qed -lemma sigma_algebra_cong: - fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme" - assumes *: "sigma_algebra M" - and cong: "space M = space M'" "sets M = sets M'" - shows "sigma_algebra M'" -using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong . - -lemma finite_Pow_additivity_sufficient: - assumes "finite (space M)" and "sets M = Pow (space M)" - and "positive \<mu>" and "additive M \<mu>" - and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>" - shows "finite_measure_space M \<mu>" -proof - - have "sigma_algebra M" - using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow]) - - have "measure_space M \<mu>" - by (rule sigma_algebra.finite_additivity_sufficient) (fact+) - thus ?thesis - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp -qed - -lemma finite_measure_spaceI: - assumes "measure_space M \<mu>" and "finite (space M)" and "sets M = Pow (space M)" - and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>" - shows "finite_measure_space M \<mu>" - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp +lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f" + unfolding simple_function_def sets_eq_Pow using finite_space by auto lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M" - unfolding measurable_def sets_eq_Pow by auto - -lemma (in finite_measure_space) simple_function_finite: "simple_function f" - unfolding simple_function_def sets_eq_Pow using finite_space by auto + by (auto intro: borel_measurable_simple_function) lemma (in finite_measure_space) positive_integral_finite_eq_setsum: "positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})" @@ -1979,10 +2078,8 @@ "positive_integral (\<lambda>x. Real (f x)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})" "positive_integral (\<lambda>x. Real (- f x)) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})" unfolding positive_integral_finite_eq_setsum by auto - show "integrable f" using finite_space finite_measure by (simp add: setsum_\<omega> integrable_def sets_eq_Pow) - show ?I using finite_measure apply (simp add: integral_def sets_eq_Pow real_of_pinfreal_setsum[symmetric] real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric])
--- a/src/HOL/Probability/Measure.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Measure.thy Fri Sep 03 13:54:39 2010 +0200 @@ -414,6 +414,19 @@ finally show ?thesis . qed +lemma (in measure_space) measure_finitely_subadditive: + assumes "finite I" "A ` I \<subseteq> sets M" + shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))" +using assms proof induct + case (insert i I) + then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN) + then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)" + using insert by (simp add: measure_subadditive) + also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))" + using insert by (auto intro!: add_left_mono) + finally show ?case . +qed simp + lemma (in measure_space) measurable_countably_subadditive: assumes "range f \<subseteq> sets M" shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))" @@ -432,9 +445,34 @@ finally show ?thesis . qed +lemma (in measure_space) measure_inter_full_set: + assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>" + assumes T: "\<mu> T = \<mu> (space M)" + shows "\<mu> (S \<inter> T) = \<mu> S" +proof (rule antisym) + show " \<mu> (S \<inter> T) \<le> \<mu> S" + using assms by (auto intro!: measure_mono) + + show "\<mu> S \<le> \<mu> (S \<inter> T)" + proof (rule ccontr) + assume contr: "\<not> ?thesis" + have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))" + unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"]) + also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)" + using assms by (auto intro!: measure_subadditive) + also have "\<dots> < \<mu> (T - S) + \<mu> S" + by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto) + also have "\<dots> = \<mu> (T \<union> S)" + using assms by (subst measure_additive) auto + also have "\<dots> \<le> \<mu> (space M)" + using assms sets_into_space by (auto intro!: measure_mono) + finally show False .. + qed +qed + lemma (in measure_space) restricted_measure_space: assumes "S \<in> sets M" - shows "measure_space (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>" + shows "measure_space (restricted_space S) \<mu>" (is "measure_space ?r \<mu>") unfolding measure_space_def measure_space_axioms_def proof safe @@ -451,6 +489,46 @@ qed qed +lemma (in measure_space) measure_space_vimage: + assumes "f \<in> measurable M M'" + and "sigma_algebra M'" + shows "measure_space M' (\<lambda>A. \<mu> (f -` A \<inter> space M))" (is "measure_space M' ?T") +proof - + interpret M': sigma_algebra M' by fact + + show ?thesis + proof + show "?T {} = 0" by simp + + show "countably_additive M' ?T" + proof (unfold countably_additive_def, safe) + fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A" + hence *: "\<And>i. f -` (A i) \<inter> space M \<in> sets M" + using `f \<in> measurable M M'` by (auto simp: measurable_def) + moreover have "(\<Union>i. f -` A i \<inter> space M) \<in> sets M" + using * by blast + moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)" + using `disjoint_family A` by (auto simp: disjoint_family_on_def) + ultimately show "(\<Sum>\<^isub>\<infinity> i. ?T (A i)) = ?T (\<Union>i. A i)" + using measure_countably_additive[OF _ **] by (auto simp: comp_def vimage_UN) + qed + qed +qed + +lemma (in measure_space) measure_space_subalgebra: + assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" + shows "measure_space (M\<lparr> sets := N \<rparr>) \<mu>" +proof - + interpret N: sigma_algebra "M\<lparr> sets := N \<rparr>" by fact + show ?thesis + proof + show "countably_additive (M\<lparr>sets := N\<rparr>) \<mu>" + using `N \<subseteq> sets M` + by (auto simp add: countably_additive_def + intro!: measure_countably_additive) + qed simp +qed + section "@{text \<sigma>}-finite Measures" locale sigma_finite_measure = measure_space + @@ -458,7 +536,7 @@ lemma (in sigma_finite_measure) restricted_sigma_finite_measure: assumes "S \<in> sets M" - shows "sigma_finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>" + shows "sigma_finite_measure (restricted_space S) \<mu>" (is "sigma_finite_measure ?r _") unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def proof safe @@ -486,6 +564,25 @@ qed qed +lemma (in sigma_finite_measure) disjoint_sigma_finite: + "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> + (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A" +proof - + obtain A :: "nat \<Rightarrow> 'a set" where + range: "range A \<subseteq> sets M" and + space: "(\<Union>i. A i) = space M" and + measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" + using sigma_finite by auto + note range' = range_disjointed_sets[OF range] range + { fix i + have "\<mu> (disjointed A i) \<le> \<mu> (A i)" + using range' disjointed_subset[of A i] by (auto intro!: measure_mono) + then have "\<mu> (disjointed A i) \<noteq> \<omega>" + using measure[of i] by auto } + with disjoint_family_disjointed UN_disjointed_eq[of A] space range' + show ?thesis by (auto intro!: exI[of _ "disjointed A"]) +qed + section "Real measure values" lemma (in measure_space) real_measure_Union: @@ -604,7 +701,7 @@ using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"]) qed -lemma (in finite_measure) finite_measure: +lemma (in finite_measure) finite_measure[simp, intro]: assumes "A \<in> sets M" shows "\<mu> A \<noteq> \<omega>" proof - @@ -619,7 +716,7 @@ lemma (in finite_measure) restricted_finite_measure: assumes "S \<in> sets M" - shows "finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>" + shows "finite_measure (restricted_space S) \<mu>" (is "finite_measure ?r _") unfolding finite_measure_def finite_measure_axioms_def proof (safe del: notI) @@ -707,6 +804,13 @@ shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s" using measure_compl[OF s, OF finite_measure, OF s] . +lemma (in finite_measure) finite_measure_inter_full_set: + assumes "S \<in> sets M" "T \<in> sets M" + assumes T: "\<mu> T = \<mu> (space M)" + shows "\<mu> (S \<inter> T) = \<mu> S" + using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms + by auto + section {* Measure preserving *} definition "measure_preserving A \<mu> B \<nu> = @@ -817,10 +921,51 @@ and sets_eq_Pow: "sets M = Pow (space M)" and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>" +lemma (in finite_measure_space) sets_image_space_eq_Pow: + "sets (image_space X) = Pow (space (image_space X))" +proof safe + fix x S assume "S \<in> sets (image_space X)" "x \<in> S" + then show "x \<in> space (image_space X)" + using sets_into_space by (auto intro!: imageI simp: image_space_def) +next + fix S assume "S \<subseteq> space (image_space X)" + then obtain S' where "S = X`S'" "S'\<in>sets M" + by (auto simp: subset_image_iff sets_eq_Pow image_space_def) + then show "S \<in> sets (image_space X)" + by (auto simp: image_space_def) +qed + lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)" using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"] by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) +lemma finite_measure_spaceI: + assumes "finite (space M)" "sets M = Pow(space M)" and space: "\<mu> (space M) \<noteq> \<omega>" + and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B" + and "\<mu> {} = 0" + shows "finite_measure_space M \<mu>" + unfolding finite_measure_space_def finite_measure_space_axioms_def +proof (safe del: notI) + show "measure_space M \<mu>" + proof (rule sigma_algebra.finite_additivity_sufficient) + show "sigma_algebra M" + apply (rule sigma_algebra_cong) + apply (rule sigma_algebra_Pow[of "space M"]) + using assms by simp_all + show "finite (space M)" by fact + show "positive \<mu>" unfolding positive_def by fact + show "additive M \<mu>" unfolding additive_def using assms by simp + qed + show "finite (space M)" by fact + { fix A x assume "A \<in> sets M" "x \<in> A" then show "x \<in> space M" + using assms by auto } + { fix A assume "A \<subseteq> space M" then show "A \<in> sets M" + using assms by auto } + { fix x assume *: "x \<in> space M" + with add[of "{x}" "space M - {x}"] space + show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) } +qed + sublocale finite_measure_space < finite_measure proof show "\<mu> (space M) \<noteq> \<omega>" @@ -828,6 +973,22 @@ using finite_space finite_single_measure by auto qed +lemma finite_measure_space_iff: + "finite_measure_space M \<mu> \<longleftrightarrow> + finite (space M) \<and> sets M = Pow(space M) \<and> \<mu> (space M) \<noteq> \<omega> \<and> \<mu> {} = 0 \<and> + (\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B)" + (is "_ = ?rhs") +proof (intro iffI) + assume "finite_measure_space M \<mu>" + then interpret finite_measure_space M \<mu> . + show ?rhs + using finite_space sets_eq_Pow measure_additive empty_measure finite_measure + by auto +next + assume ?rhs then show "finite_measure_space M \<mu>" + by (auto intro!: finite_measure_spaceI) +qed + lemma psuminf_cmult_indicator: assumes "disjoint_family A" "x \<in> A i" shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
--- a/src/HOL/Probability/Positive_Infinite_Real.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Positive_Infinite_Real.thy Fri Sep 03 13:54:39 2010 +0200 @@ -411,6 +411,10 @@ lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>" by (cases x) auto +lemma pinfreal_0_less_mult_iff[simp]: + fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y" + by (cases x, cases y) (auto simp: zero_less_mult_iff) + subsection {* @{text "x - y"} on @{typ pinfreal} *} instantiation pinfreal :: minus
--- a/src/HOL/Probability/Probability_Space.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Probability_Space.thy Fri Sep 03 13:54:39 2010 +0200 @@ -1,39 +1,7 @@ theory Probability_Space -imports Lebesgue_Integration +imports Lebesgue_Integration Radon_Nikodym begin -lemma (in measure_space) measure_inter_full_set: - assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>" - assumes T: "\<mu> T = \<mu> (space M)" - shows "\<mu> (S \<inter> T) = \<mu> S" -proof (rule antisym) - show " \<mu> (S \<inter> T) \<le> \<mu> S" - using assms by (auto intro!: measure_mono) - - show "\<mu> S \<le> \<mu> (S \<inter> T)" - proof (rule ccontr) - assume contr: "\<not> ?thesis" - have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))" - unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"]) - also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)" - using assms by (auto intro!: measure_subadditive) - also have "\<dots> < \<mu> (T - S) + \<mu> S" - by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto) - also have "\<dots> = \<mu> (T \<union> S)" - using assms by (subst measure_additive) auto - also have "\<dots> \<le> \<mu> (space M)" - using assms sets_into_space by (auto intro!: measure_mono) - finally show False .. - qed -qed - -lemma (in finite_measure) finite_measure_inter_full_set: - assumes "S \<in> sets M" "T \<in> sets M" - assumes T: "\<mu> T = \<mu> (space M)" - shows "\<mu> (S \<inter> T) = \<mu> S" - using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms - by auto - locale prob_space = measure_space + assumes measure_space_1: "\<mu> (space M) = 1" @@ -63,6 +31,19 @@ abbreviation "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))" +lemma (in prob_space) distribution_cong: + assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x" + shows "distribution X = distribution Y" + unfolding distribution_def expand_fun_eq + using assms by (auto intro!: arg_cong[where f="\<mu>"]) + +lemma (in prob_space) joint_distribution_cong: + assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x" + assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x" + shows "joint_distribution X Y = joint_distribution X' Y'" + unfolding distribution_def expand_fun_eq + using assms by (auto intro!: arg_cong[where f="\<mu>"]) + lemma prob_space: "prob (space M) = 1" unfolding measure_space_1 by simp @@ -75,10 +56,6 @@ finally show ?thesis . qed -lemma measure_finite[simp, intro]: - assumes "A \<in> events" shows "\<mu> A \<noteq> \<omega>" - using measure_le_1[OF assms] by auto - lemma prob_compl: assumes "A \<in> events" shows "prob (space M - A) = 1 - prob A" @@ -197,35 +174,17 @@ qed lemma distribution_prob_space: - fixes S :: "('c, 'd) algebra_scheme" - assumes "sigma_algebra S" "random_variable S X" + assumes S: "sigma_algebra S" "random_variable S X" shows "prob_space S (distribution X)" proof - - interpret S: sigma_algebra S by fact + interpret S: measure_space S "distribution X" + using measure_space_vimage[OF S(2,1)] unfolding distribution_def . show ?thesis proof - show "distribution X {} = 0" unfolding distribution_def by simp have "X -` space S \<inter> space M = space M" using `random_variable S X` by (auto simp: measurable_def) - then show "distribution X (space S) = 1" using measure_space_1 by (simp add: distribution_def) - - show "countably_additive S (distribution X)" - proof (unfold countably_additive_def, safe) - fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets S" "disjoint_family A" - hence *: "\<And>i. X -` A i \<inter> space M \<in> sets M" - using `random_variable S X` by (auto simp: measurable_def) - moreover hence "\<And>i. \<mu> (X -` A i \<inter> space M) \<noteq> \<omega>" - using finite_measure by auto - moreover have "(\<Union>i. X -` A i \<inter> space M) \<in> sets M" - using * by blast - moreover hence "\<mu> (\<Union>i. X -` A i \<inter> space M) \<noteq> \<omega>" - using finite_measure by auto - moreover have **: "disjoint_family (\<lambda>i. X -` A i \<inter> space M)" - using `disjoint_family A` by (auto simp: disjoint_family_on_def) - ultimately show "(\<Sum>\<^isub>\<infinity> i. distribution X (A i)) = distribution X (\<Union>i. A i)" - using measure_countably_additive[OF _ **] - by (auto simp: distribution_def Real_real comp_def vimage_UN) - qed + then show "distribution X (space S) = 1" + using measure_space_1 by (simp add: distribution_def) qed qed @@ -379,39 +338,246 @@ joint_distribution_restriction_snd[of X Y "{(x, y)}"] by auto +lemma (in finite_prob_space) distribution_mono: + assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" + shows "distribution X x \<le> distribution Y y" + unfolding distribution_def + using assms by (auto simp: sets_eq_Pow intro!: measure_mono) + +lemma (in finite_prob_space) distribution_mono_gt_0: + assumes gt_0: "0 < distribution X x" + assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y" + shows "0 < distribution Y y" + by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) + +lemma (in finite_prob_space) sum_over_space_distrib: + "(\<Sum>x\<in>X`space M. distribution X {x}) = 1" + unfolding distribution_def measure_space_1[symmetric] using finite_space + by (subst measure_finitely_additive'') + (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>]) + +lemma (in finite_prob_space) sum_over_space_real_distribution: + "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1" + unfolding distribution_def prob_space[symmetric] using finite_space + by (subst real_finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) finite_sum_over_space_eq_1: + "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1" + using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow) + +lemma (in finite_prob_space) distribution_finite: + "distribution X A \<noteq> \<omega>" + using finite_measure[of "X -` A \<inter> space M"] + unfolding distribution_def sets_eq_Pow by auto + +lemma (in finite_prob_space) real_distribution_gt_0[simp]: + "0 < real (distribution Y y) \<longleftrightarrow> 0 < distribution Y y" + using assms by (auto intro!: real_pinfreal_pos distribution_finite) + +lemma (in finite_prob_space) real_distribution_mult_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y * distribution X x)" + unfolding real_of_pinfreal_mult[symmetric] + using assms by (auto intro!: mult_pos_pos) + +lemma (in finite_prob_space) real_distribution_divide_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y / distribution X x)" + unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] + using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) + +lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y * inverse (distribution X x))" + unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] + using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) + +lemma (in prob_space) distribution_remove_const: + shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}" + and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}" + and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" + and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" + and "distribution (\<lambda>x. ()) {()} = 1" + unfolding measure_space_1[symmetric] + by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def) + +lemma (in finite_prob_space) setsum_distribution_gen: + assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M" + and "inj_on f (X`space M)" + shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}" + unfolding distribution_def assms + using finite_space assms + by (subst measure_finitely_additive'') + (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def + intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) setsum_distribution: + "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" + "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" + "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" + "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" + "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" + by (auto intro!: inj_onI setsum_distribution_gen) + +lemma (in finite_prob_space) setsum_real_distribution_gen: + assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M" + and "inj_on f (X`space M)" + shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})" + unfolding distribution_def assms + using finite_space assms + by (subst real_finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def + intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) setsum_real_distribution: + "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})" + "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})" + "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})" + "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})" + "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})" + by (auto intro!: inj_onI setsum_real_distribution_gen) + +lemma (in finite_prob_space) real_distribution_order: + shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})" + and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})" + and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})" + and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})" + and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" + and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] + using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] + by auto + +lemma (in prob_space) joint_distribution_remove[simp]: + "joint_distribution X X {(x, x)} = distribution X {x}" + unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"]) + +lemma (in finite_prob_space) distribution_1: + "distribution X A \<le> 1" + unfolding distribution_def measure_space_1[symmetric] + by (auto intro!: measure_mono simp: sets_eq_Pow) + +lemma (in finite_prob_space) real_distribution_1: + "real (distribution X A) \<le> 1" + unfolding real_pinfreal_1[symmetric] + by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp + +lemma (in finite_prob_space) uniform_prob: + assumes "x \<in> space M" + assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}" + shows "prob {x} = 1 / real (card (space M))" +proof - + have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}" + using assms(2)[OF _ `x \<in> space M`] by blast + have "1 = prob (space M)" + using prob_space by auto + also have "\<dots> = (\<Sum> x \<in> space M. prob {x})" + using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified] + sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] + finite_space unfolding disjoint_family_on_def prob_space[symmetric] + by (auto simp add:setsum_restrict_set) + also have "\<dots> = (\<Sum> y \<in> space M. prob {x})" + using prob_x by auto + also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp + finally have one: "1 = real (card (space M)) * prob {x}" + using real_eq_of_nat by auto + hence two: "real (card (space M)) \<noteq> 0" by fastsimp + from one have three: "prob {x} \<noteq> 0" by fastsimp + thus ?thesis using one two three divide_cancel_right + by (auto simp:field_simps) +qed + +lemma (in prob_space) prob_space_subalgebra: + assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" + shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>" +proof - + interpret N: measure_space "M\<lparr> sets := N \<rparr>" \<mu> + using measure_space_subalgebra[OF assms] . + show ?thesis + proof qed (simp add: measure_space_1) +qed + +lemma (in prob_space) prob_space_of_restricted_space: + assumes "\<mu> A \<noteq> 0" "\<mu> A \<noteq> \<omega>" "A \<in> sets M" + shows "prob_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)" + unfolding prob_space_def prob_space_axioms_def +proof + show "\<mu> (space (restricted_space A)) / \<mu> A = 1" + using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pinfreal_noteq_omega_Ex) + have *: "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute) + interpret A: measure_space "restricted_space A" \<mu> + using `A \<in> sets M` by (rule restricted_measure_space) + show "measure_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)" + proof + show "\<mu> {} / \<mu> A = 0" by auto + show "countably_additive (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)" + unfolding countably_additive_def psuminf_cmult_right * + using A.measure_countably_additive by auto + qed +qed + +lemma finite_prob_spaceI: + assumes "finite (space M)" "sets M = Pow(space M)" "\<mu> (space M) = 1" "\<mu> {} = 0" + and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B" + shows "finite_prob_space M \<mu>" + unfolding finite_prob_space_eq +proof + show "finite_measure_space M \<mu>" using assms + by (auto intro!: finite_measure_spaceI) + show "\<mu> (space M) = 1" by fact +qed + +lemma (in finite_prob_space) finite_measure_space: + fixes X :: "'a \<Rightarrow> 'x" + shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" + (is "finite_measure_space ?S _") +proof (rule finite_measure_spaceI, simp_all) + show "finite (X ` space M)" using finite_space by simp +next + fix A B :: "'x set" assume "A \<inter> B = {}" + then show "distribution X (A \<union> B) = distribution X A + distribution X B" + unfolding distribution_def + by (subst measure_additive) + (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow) +qed + +lemma (in finite_prob_space) finite_prob_space_of_images: + "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" + by (simp add: finite_prob_space_eq finite_measure_space) + +lemma (in prob_space) joint_distribution_commute: + "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)" + unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) + +lemma (in finite_prob_space) real_distribution_order': + shows "real (distribution X {x}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" + and "real (distribution Y {y}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0" + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] + using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] + by auto + lemma (in finite_prob_space) finite_product_measure_space: + fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y" assumes "finite s1" "finite s2" shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)" (is "finite_measure_space ?M ?D") -proof (rule finite_Pow_additivity_sufficient) - show "positive ?D" - unfolding positive_def using assms sets_eq_Pow - by (simp add: distribution_def) - - show "additive ?M ?D" unfolding additive_def - proof safe - fix x y - have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto - have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto - assume "x \<inter> y = {}" - hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}" - by auto - from additive[unfolded additive_def, rule_format, OF A B] this - finite_measure[OF A] finite_measure[OF B] - show "?D (x \<union> y) = ?D x + ?D y" - apply (simp add: distribution_def) - apply (subst Int_Un_distrib2) - by (auto simp: real_of_pinfreal_add) - qed - - show "finite (space ?M)" +proof (rule finite_measure_spaceI, simp_all) + show "finite (s1 \<times> s2)" using assms by auto - - show "sets ?M = Pow (space ?M)" - by simp - - { fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>" - unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } + show "joint_distribution X Y (s1\<times>s2) \<noteq> \<omega>" + using distribution_finite . +next + fix A B :: "('x*'y) set" assume "A \<inter> B = {}" + then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B" + unfolding distribution_def + by (subst measure_additive) + (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow) qed lemma (in finite_prob_space) finite_product_measure_space_of_images: @@ -420,47 +586,133 @@ (joint_distribution X Y)" using finite_space by (auto intro!: finite_product_measure_space) -lemma (in finite_prob_space) finite_measure_space: - shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" - (is "finite_measure_space ?S _") -proof (rule finite_Pow_additivity_sufficient, simp_all) - show "finite (X ` space M)" using finite_space by simp +section "Conditional Expectation and Probability" - show "positive (distribution X)" - unfolding distribution_def positive_def using sets_eq_Pow by auto +lemma (in prob_space) conditional_expectation_exists: + fixes X :: "'a \<Rightarrow> pinfreal" + assumes borel: "X \<in> borel_measurable M" + and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" + shows "\<exists>Y\<in>borel_measurable (M\<lparr> sets := N \<rparr>). \<forall>C\<in>N. + positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)" +proof - + interpret P: prob_space "M\<lparr> sets := N \<rparr>" \<mu> + using prob_space_subalgebra[OF N_subalgebra] . + + let "?f A" = "\<lambda>x. X x * indicator A x" + let "?Q A" = "positive_integral (?f A)" + + from measure_space_density[OF borel] + have Q: "measure_space (M\<lparr> sets := N \<rparr>) ?Q" + by (rule measure_space.measure_space_subalgebra[OF _ N_subalgebra]) + then interpret Q: measure_space "M\<lparr> sets := N \<rparr>" ?Q . - show "additive ?S (distribution X)" unfolding additive_def distribution_def - proof (simp, safe) - fix x y - have x: "(X -` x) \<inter> space M \<in> sets M" - and y: "(X -` y) \<inter> space M \<in> sets M" using sets_eq_Pow by auto - assume "x \<inter> y = {}" - hence "X -` x \<inter> space M \<inter> (X -` y \<inter> space M) = {}" by auto - from additive[unfolded additive_def, rule_format, OF x y] this - finite_measure[OF x] finite_measure[OF y] - have "\<mu> (((X -` x) \<union> (X -` y)) \<inter> space M) = - \<mu> ((X -` x) \<inter> space M) + \<mu> ((X -` y) \<inter> space M)" - by (subst Int_Un_distrib2) auto - thus "\<mu> ((X -` x \<union> X -` y) \<inter> space M) = \<mu> (X -` x \<inter> space M) + \<mu> (X -` y \<inter> space M)" - by auto + have "P.absolutely_continuous ?Q" + unfolding P.absolutely_continuous_def + proof (safe, simp) + fix A assume "A \<in> N" "\<mu> A = 0" + moreover then have f_borel: "?f A \<in> borel_measurable M" + using borel N_subalgebra by (auto intro: borel_measurable_indicator) + moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A" + by (auto simp: indicator_def) + moreover have "\<mu> \<dots> \<le> \<mu> A" + using `A \<in> N` N_subalgebra f_borel + by (auto intro!: measure_mono Int[of _ A] measurable_sets) + ultimately show "?Q A = 0" + by (simp add: positive_integral_0_iff) qed - - { fix x assume "x \<in> X ` space M" thus "distribution X {x} \<noteq> \<omega>" - unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } + from P.Radon_Nikodym[OF Q this] + obtain Y where Y: "Y \<in> borel_measurable (M\<lparr>sets := N\<rparr>)" + "\<And>A. A \<in> sets (M\<lparr>sets:=N\<rparr>) \<Longrightarrow> ?Q A = P.positive_integral (\<lambda>x. Y x * indicator A x)" + by blast + with N_subalgebra show ?thesis + by (auto intro!: bexI[OF _ Y(1)]) qed -lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)" - by (simp add: finite_prob_space_eq finite_measure_space) +definition (in prob_space) + "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable (M\<lparr>sets:=N\<rparr>) + \<and> (\<forall>C\<in>N. positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)))" + +abbreviation (in prob_space) + "conditional_prob N A \<equiv> conditional_expectation N (indicator A)" + +lemma (in prob_space) + fixes X :: "'a \<Rightarrow> pinfreal" + assumes borel: "X \<in> borel_measurable M" + and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)" + shows borel_measurable_conditional_expectation: + "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)" + and conditional_expectation: "\<And>C. C \<in> N \<Longrightarrow> + positive_integral (\<lambda>x. conditional_expectation N X x * indicator C x) = + positive_integral (\<lambda>x. X x * indicator C x)" + (is "\<And>C. C \<in> N \<Longrightarrow> ?eq C") +proof - + note CE = conditional_expectation_exists[OF assms, unfolded Bex_def] + then show "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)" + unfolding conditional_expectation_def by (rule someI2_ex) blast + + from CE show "\<And>C. C\<in>N \<Longrightarrow> ?eq C" + unfolding conditional_expectation_def by (rule someI2_ex) blast +qed + +lemma (in sigma_algebra) factorize_measurable_function: + fixes Z :: "'a \<Rightarrow> pinfreal" and Y :: "'a \<Rightarrow> 'c" + assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M" + shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) + \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))" +proof safe + interpret M': sigma_algebra M' by fact + have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto + from M'.sigma_algebra_vimage[OF this] + interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . -lemma (in finite_prob_space) finite_product_prob_space_of_images: - "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr> - (joint_distribution X Y)" - (is "finite_prob_space ?S _") -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) - have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto - thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1" - by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) + { fix g :: "'c \<Rightarrow> pinfreal" assume "g \<in> borel_measurable M'" + with M'.measurable_vimage_algebra[OF Y] + have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" + by (rule measurable_comp) + moreover assume "\<forall>x\<in>space M. Z x = g (Y x)" + then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow> + g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)" + by (auto intro!: measurable_cong) + ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" + by simp } + + assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)" + from va.borel_measurable_implies_simple_function_sequence[OF this] + obtain f where f: "\<And>i. va.simple_function (f i)" and "f \<up> Z" by blast + + have "\<forall>i. \<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" + proof + fix i + from f[of i] have "finite (f i`space M)" and B_ex: + "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M" + unfolding va.simple_function_def by auto + from B_ex[THEN bchoice] guess B .. note B = this + + let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x" + + show "\<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))" + proof (intro exI[of _ ?g] conjI ballI) + show "M'.simple_function ?g" using B by auto + + fix x assume "x \<in> space M" + then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pinfreal)" + unfolding indicator_def using B by auto + then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i] + by (subst va.simple_function_indicator_representation) auto + qed + qed + from choice[OF this] guess g .. note g = this + + show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)" + proof (intro ballI bexI) + show "(SUP i. g i) \<in> borel_measurable M'" + using g by (auto intro: M'.borel_measurable_simple_function) + fix x assume "x \<in> space M" + have "Z x = (SUP i. f i) x" using `f \<up> Z` unfolding isoton_def by simp + also have "\<dots> = (SUP i. g i) (Y x)" unfolding SUPR_fun_expand + using g `x \<in> space M` by simp + finally show "Z x = (SUP i. g i) (Y x)" . + qed qed end
--- a/src/HOL/Probability/Product_Measure.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Product_Measure.thy Fri Sep 03 13:54:39 2010 +0200 @@ -2,14 +2,412 @@ imports Lebesgue_Integration begin +definition "dynkin M \<longleftrightarrow> + space M \<in> sets M \<and> + (\<forall> A \<in> sets M. A \<subseteq> space M) \<and> + (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M) \<and> + (\<forall>A. disjoint_family A \<and> range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)" + +lemma dynkinI: + assumes "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M" + assumes "space M \<in> sets M" and "\<forall> b \<in> sets M. \<forall> a \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M" + assumes "\<And> a. (\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}) + \<Longrightarrow> (\<And> i :: nat. a i \<in> sets M) \<Longrightarrow> UNION UNIV a \<in> sets M" + shows "dynkin M" +using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) + +lemma dynkin_subset: + assumes "dynkin M" + shows "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M" +using assms unfolding dynkin_def by auto + +lemma dynkin_space: + assumes "dynkin M" + shows "space M \<in> sets M" +using assms unfolding dynkin_def by auto + +lemma dynkin_diff: + assumes "dynkin M" + shows "\<And> a b. \<lbrakk> a \<in> sets M ; b \<in> sets M ; a \<subseteq> b \<rbrakk> \<Longrightarrow> b - a \<in> sets M" +using assms unfolding dynkin_def by auto + +lemma dynkin_empty: + assumes "dynkin M" + shows "{} \<in> sets M" +using dynkin_diff[OF assms dynkin_space[OF assms] dynkin_space[OF assms]] by auto + +lemma dynkin_UN: + assumes "dynkin M" + assumes "\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" + assumes "\<And> i :: nat. a i \<in> sets M" + shows "UNION UNIV a \<in> sets M" +using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) + +definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)" + +lemma dynkin_trivial: + shows "dynkin \<lparr> space = A, sets = Pow A \<rparr>" +by (rule dynkinI) auto + +lemma dynkin_lemma: + fixes D :: "'a algebra" + assumes stab: "Int_stable E" + and spac: "space E = space D" + and subsED: "sets E \<subseteq> sets D" + and subsDE: "sets D \<subseteq> sigma_sets (space E) (sets E)" + and dyn: "dynkin D" + shows "sigma (space E) (sets E) = D" +proof - + def sets_\<delta>E == "\<Inter> {sets d | d :: 'a algebra. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}" + def \<delta>E == "\<lparr> space = space E, sets = sets_\<delta>E \<rparr>" + have "\<lparr> space = space E, sets = Pow (space E) \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}" + using dynkin_trivial spac subsED dynkin_subset[OF dyn] by fastsimp + hence not_empty: "{sets (d :: 'a algebra) | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d} \<noteq> {}" + using exI[of "\<lambda> x. space x = space E \<and> dynkin x \<and> sets E \<subseteq> sets x" "\<lparr> space = space E, sets = Pow (space E) \<rparr>", simplified] + by auto + have \<delta>E_D: "sets_\<delta>E \<subseteq> sets D" + unfolding sets_\<delta>E_def using assms by auto + have \<delta>ynkin: "dynkin \<delta>E" + proof (rule dynkinI, safe) + fix A x assume asm: "A \<in> sets \<delta>E" "x \<in> A" + { fix d :: "('a, 'b) algebra_scheme" assume "A \<in> sets d" "dynkin d \<and> space d = space E" + hence "A \<subseteq> space d" using dynkin_subset by auto } + show "x \<in> space \<delta>E" using asm unfolding \<delta>E_def sets_\<delta>E_def using not_empty + by simp (metis dynkin_subset in_mono mem_def) + next + show "space \<delta>E \<in> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def + using dynkin_space by fastsimp + next + fix a b assume "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" "a \<subseteq> b" + thus "b - a \<in> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def by (auto intro:dynkin_diff) + next + fix a assume asm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> sets \<delta>E" + thus "UNION UNIV a \<in> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def apply (auto intro!:dynkin_UN[OF _ asm(1)]) + by blast + qed + + def Dy == "\<lambda> d. {A | A. A \<in> sets_\<delta>E \<and> A \<inter> d \<in> sets_\<delta>E}" + { fix d assume dasm: "d \<in> sets_\<delta>E" + have "dynkin \<lparr> space = space E, sets = Dy d \<rparr>" + proof (rule dynkinI, safe, simp_all) + fix A x assume "A \<in> Dy d" "x \<in> A" + thus "x \<in> space E" unfolding Dy_def sets_\<delta>E_def using not_empty + by simp (metis dynkin_subset in_mono mem_def) + next + show "space E \<in> Dy d" + unfolding Dy_def \<delta>E_def sets_\<delta>E_def + proof auto + fix d assume asm: "dynkin d" "space d = space E" "sets E \<subseteq> sets d" + hence "space d \<in> sets d" using dynkin_space[OF asm(1)] by auto + thus "space E \<in> sets d" using asm by auto + next + fix da :: "'a algebra" assume asm: "dynkin da" "space da = space E" "sets E \<subseteq> sets da" + have d: "d = space E \<inter> d" + using dasm dynkin_subset[OF asm(1)] asm(2) dynkin_subset[OF \<delta>ynkin] + unfolding \<delta>E_def by auto + hence "space E \<inter> d \<in> sets \<delta>E" unfolding \<delta>E_def + using dasm by auto + have "sets \<delta>E \<subseteq> sets da" unfolding \<delta>E_def sets_\<delta>E_def using asm + by auto + thus "space E \<inter> d \<in> sets da" using dasm asm d dynkin_subset[OF \<delta>ynkin] + unfolding \<delta>E_def by auto + qed + next + fix a b assume absm: "a \<in> Dy d" "b \<in> Dy d" "a \<subseteq> b" + hence "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" + unfolding Dy_def \<delta>E_def by auto + hence *: "b - a \<in> sets \<delta>E" + using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto + have "a \<inter> d \<in> sets \<delta>E" "b \<inter> d \<in> sets \<delta>E" + using absm unfolding Dy_def \<delta>E_def by auto + hence "(b \<inter> d) - (a \<inter> d) \<in> sets \<delta>E" + using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto + hence **: "(b - a) \<inter> d \<in> sets \<delta>E" by (auto simp add:Diff_Int_distrib2) + thus "b - a \<in> Dy d" + using * ** unfolding Dy_def \<delta>E_def by auto + next + fix a assume aasm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> Dy d" + hence "\<And> i. a i \<in> sets \<delta>E" + unfolding Dy_def \<delta>E_def by auto + from dynkin_UN[OF \<delta>ynkin aasm(1) this] + have *: "UNION UNIV a \<in> sets \<delta>E" by auto + from aasm + have aE: "\<forall> i. a i \<inter> d \<in> sets \<delta>E" + unfolding Dy_def \<delta>E_def by auto + from aasm + have "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> (a i \<inter> d) \<inter> (a j \<inter> d) = {}" by auto + from dynkin_UN[OF \<delta>ynkin this] + have "UNION UNIV (\<lambda> i. a i \<inter> d) \<in> sets \<delta>E" + using aE by auto + hence **: "UNION UNIV a \<inter> d \<in> sets \<delta>E" by auto + from * ** show "UNION UNIV a \<in> Dy d" unfolding Dy_def \<delta>E_def by auto + qed } note Dy_nkin = this + have E_\<delta>E: "sets E \<subseteq> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def by auto + { fix d assume dasm: "d \<in> sets \<delta>E" + { fix e assume easm: "e \<in> sets E" + hence deasm: "e \<in> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def by auto + have subset: "Dy e \<subseteq> sets \<delta>E" + unfolding Dy_def \<delta>E_def by auto + { fix e' assume e'asm: "e' \<in> sets E" + have "e' \<inter> e \<in> sets E" + using easm e'asm stab unfolding Int_stable_def by auto + hence "e' \<inter> e \<in> sets \<delta>E" + unfolding \<delta>E_def sets_\<delta>E_def by auto + hence "e' \<in> Dy e" using e'asm unfolding Dy_def \<delta>E_def sets_\<delta>E_def by auto } + hence E_Dy: "sets E \<subseteq> Dy e" by auto + have "\<lparr> space = space E, sets = Dy e \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}" + using Dy_nkin[OF deasm[unfolded \<delta>E_def, simplified]] E_\<delta>E E_Dy by auto + hence "sets_\<delta>E \<subseteq> Dy e" + unfolding sets_\<delta>E_def by auto (metis E_Dy simps(1) simps(2) spac) + hence "sets \<delta>E = Dy e" using subset unfolding \<delta>E_def by auto + hence "d \<inter> e \<in> sets \<delta>E" + using dasm easm deasm unfolding Dy_def \<delta>E_def by auto + hence "e \<in> Dy d" using deasm + unfolding Dy_def \<delta>E_def + by (auto simp add:Int_commute) } + hence "sets E \<subseteq> Dy d" by auto + hence "sets \<delta>E \<subseteq> Dy d" using Dy_nkin[OF dasm[unfolded \<delta>E_def, simplified]] + unfolding \<delta>E_def sets_\<delta>E_def + by auto (metis `sets E <= Dy d` simps(1) simps(2) spac) + hence *: "sets \<delta>E = Dy d" + unfolding Dy_def \<delta>E_def by auto + fix a assume aasm: "a \<in> sets \<delta>E" + hence "a \<inter> d \<in> sets \<delta>E" + using * dasm unfolding Dy_def \<delta>E_def by auto } note \<delta>E_stab = this + { fix A :: "nat \<Rightarrow> 'a set" assume Asm: "range A \<subseteq> sets \<delta>E" "\<And>A. A \<in> sets \<delta>E \<Longrightarrow> A \<subseteq> space \<delta>E" + "\<And>a. a \<in> sets \<delta>E \<Longrightarrow> space \<delta>E - a \<in> sets \<delta>E" + "{} \<in> sets \<delta>E" "space \<delta>E \<in> sets \<delta>E" + let "?A i" = "A i \<inter> (\<Inter> j \<in> {..< i}. space \<delta>E - A j)" + { fix i :: nat + have *: "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<inter> space \<delta>E \<in> sets \<delta>E" + apply (induct i) + using lessThan_Suc Asm \<delta>E_stab apply fastsimp + apply (subst lessThan_Suc) + apply (subst INT_insert) + apply (subst Int_assoc) + apply (subst \<delta>E_stab) + using lessThan_Suc Asm \<delta>E_stab Asm + apply (fastsimp simp add:Int_assoc dynkin_diff[OF \<delta>ynkin]) + prefer 2 apply simp + apply (rule dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]]) + using Asm by auto + have **: "\<And> i. A i \<subseteq> space \<delta>E" using Asm by auto + have "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<subseteq> space \<delta>E \<or> (\<Inter> j \<in> {..< i}. A j) = UNIV \<and> i = 0" + apply (cases i) + using Asm ** dynkin_subset[OF \<delta>ynkin, of "A (i - 1)"] + by auto + hence Aisets: "?A i \<in> sets \<delta>E" + apply (cases i) + using Asm * apply fastsimp + apply (rule \<delta>E_stab) + using Asm * ** + by (auto simp add:Int_absorb2) + have "?A i = disjointed A i" unfolding disjointed_def + atLeast0LessThan using Asm by auto + hence "?A i = disjointed A i" "?A i \<in> sets \<delta>E" + using Aisets by auto + } note Ai = this + from dynkin_UN[OF \<delta>ynkin _ this(2)] this disjoint_family_disjointed[of A] + have "(\<Union> i. ?A i) \<in> sets \<delta>E" + by (auto simp add:disjoint_family_on_def disjointed_def) + hence "(\<Union> i. A i) \<in> sets \<delta>E" + using Ai(1) UN_disjointed_eq[of A] by auto } note \<delta>E_UN = this + { fix a b assume asm: "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" + let ?ab = "\<lambda> i. if (i::nat) = 0 then a else if i = 1 then b else {}" + have *: "(\<Union> i. ?ab i) \<in> sets \<delta>E" + apply (rule \<delta>E_UN) + using asm \<delta>E_UN dynkin_empty[OF \<delta>ynkin] + dynkin_subset[OF \<delta>ynkin] + dynkin_space[OF \<delta>ynkin] + dynkin_diff[OF \<delta>ynkin] by auto + have "(\<Union> i. ?ab i) = a \<union> b" apply auto + apply (case_tac "i = 0") + apply auto + apply (case_tac "i = 1") + by auto + hence "a \<union> b \<in> sets \<delta>E" using * by auto} note \<delta>E_Un = this + have "sigma_algebra \<delta>E" + apply unfold_locales + using dynkin_subset[OF \<delta>ynkin] + using dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]] + using dynkin_diff[OF \<delta>ynkin, of "space \<delta>E" "space \<delta>E", OF dynkin_space[OF \<delta>ynkin] dynkin_space[OF \<delta>ynkin]] + using dynkin_space[OF \<delta>ynkin] + using \<delta>E_UN \<delta>E_Un + by auto + from sigma_algebra.sigma_subset[OF this E_\<delta>E] \<delta>E_D subsDE spac + show ?thesis by (auto simp add:\<delta>E_def sigma_def) +qed + +lemma measure_eq: + assumes fin: "\<mu> (space M) = \<nu> (space M)" "\<nu> (space M) < \<omega>" + assumes E: "M = sigma (space E) (sets E)" "Int_stable E" + assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e" + assumes ms: "measure_space M \<mu>" "measure_space M \<nu>" + assumes A: "A \<in> sets M" + shows "\<mu> A = \<nu> A" +proof - + interpret M: measure_space M \<mu> + using ms by simp + interpret M': measure_space M \<nu> + using ms by simp + + let ?D_sets = "{A. A \<in> sets M \<and> \<mu> A = \<nu> A}" + have \<delta>: "dynkin \<lparr> space = space M , sets = ?D_sets \<rparr>" + proof (rule dynkinI, safe, simp_all) + fix A x assume "A \<in> sets M \<and> \<mu> A = \<nu> A" "x \<in> A" + thus "x \<in> space M" using assms M.sets_into_space by auto + next + show "\<mu> (space M) = \<nu> (space M)" + using fin by auto + next + fix a b + assume asm: "a \<in> sets M \<and> \<mu> a = \<nu> a" + "b \<in> sets M \<and> \<mu> b = \<nu> b" "a \<subseteq> b" + hence "a \<subseteq> space M" + using M.sets_into_space by auto + from M.measure_mono[OF this] + have "\<mu> a \<le> \<mu> (space M)" + using asm by auto + hence afin: "\<mu> a < \<omega>" + using fin by auto + have *: "b = b - a \<union> a" using asm by auto + have **: "(b - a) \<inter> a = {}" using asm by auto + have iv: "\<mu> (b - a) + \<mu> a = \<mu> b" + using M.measure_additive[of "b - a" a] + conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** + by auto + have v: "\<nu> (b - a) + \<nu> a = \<nu> b" + using M'.measure_additive[of "b - a" a] + conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** + by auto + from iv v have "\<mu> (b - a) = \<nu> (b - a)" using asm afin + pinfreal_add_cancel_right[of "\<mu> (b - a)" "\<nu> a" "\<nu> (b - a)"] + by auto + thus "b - a \<in> sets M \<and> \<mu> (b - a) = \<nu> (b - a)" + using asm by auto + next + fix a assume "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" + "\<And>i. a i \<in> sets M \<and> \<mu> (a i) = \<nu> (a i)" + thus "(\<Union>x. a x) \<in> sets M \<and> \<mu> (\<Union>x. a x) = \<nu> (\<Union>x. a x)" + using M.measure_countably_additive + M'.measure_countably_additive + M.countable_UN + apply (auto simp add:disjoint_family_on_def image_def) + apply (subst M.measure_countably_additive[symmetric]) + apply (auto simp add:disjoint_family_on_def) + apply (subst M'.measure_countably_additive[symmetric]) + by (auto simp add:disjoint_family_on_def) + qed + have *: "sets E \<subseteq> ?D_sets" + using eq E sigma_sets.Basic[of _ "sets E"] + by (auto simp add:sigma_def) + have **: "?D_sets \<subseteq> sets M" by auto + have "M = \<lparr> space = space M , sets = ?D_sets \<rparr>" + unfolding E(1) + apply (rule dynkin_lemma[OF E(2)]) + using eq E space_sigma \<delta> sigma_sets.Basic + by (auto simp add:sigma_def) + from subst[OF this, of "\<lambda> M. A \<in> sets M", OF A] + show ?thesis by auto +qed +(* +lemma + assumes sfin: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And> i :: nat. \<nu> (A i) < \<omega>" + assumes A: "\<And> i. \<mu> (A i) = \<nu> (A i)" "\<And> i. A i \<subseteq> A (Suc i)" + assumes E: "M = sigma (space E) (sets E)" "Int_stable E" + assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e" + assumes ms: "measure_space (M :: 'a algebra) \<mu>" "measure_space M \<nu>" + assumes B: "B \<in> sets M" + shows "\<mu> B = \<nu> B" +proof - + interpret M: measure_space M \<mu> by (rule ms) + interpret M': measure_space M \<nu> by (rule ms) + have *: "M = \<lparr> space = space M, sets = sets M \<rparr>" by auto + { fix i :: nat + have **: "M\<lparr> space := A i, sets := op \<inter> (A i) ` sets M \<rparr> = + \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>" + by auto + have mu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<mu>" + using M.restricted_measure_space[of "A i", simplified **] + sfin by auto + have nu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<nu>" + using M'.restricted_measure_space[of "A i", simplified **] + sfin by auto + let ?M = "\<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>" + have "\<mu> (A i \<inter> B) = \<nu> (A i \<inter> B)" + apply (rule measure_eq[of \<mu> ?M \<nu> "\<lparr> space = space E \<inter> A i, sets = op \<inter> (A i) ` sets E\<rparr>" "A i \<inter> B", simplified]) + using assms nu_i mu_i + apply (auto simp add:image_def) (* TODO *) sorry + show ?thesis sorry +qed +*) definition prod_sets where "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}" definition - "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))" + "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))" + +lemma + fixes M1 :: "'a algebra" and M2 :: "'b algebra" + assumes "algebra M1" "algebra M2" + shows measureable_fst[intro!, simp]: + "fst \<in> measurable (prod_measure_space M1 M2) M1" (is ?fst) + and measureable_snd[intro!, simp]: + "snd \<in> measurable (prod_measure_space M1 M2) M2" (is ?snd) +proof - + interpret M1: algebra M1 by fact + interpret M2: algebra M2 by fact + + { fix X assume "X \<in> sets M1" + then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2" + apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) + using M1.sets_into_space by force+ } + moreover + { fix X assume "X \<in> sets M2" + then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2" + apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) + using M2.sets_into_space by force+ } + ultimately show ?fst ?snd + by (force intro!: sigma_sets.Basic + simp: measurable_def prod_measure_space_def prod_sets_def sets_sigma)+ +qed + +lemma (in sigma_algebra) measureable_prod: + fixes M1 :: "'a algebra" and M2 :: "'b algebra" + assumes "algebra M1" "algebra M2" + shows "f \<in> measurable M (prod_measure_space M1 M2) \<longleftrightarrow> + (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2" +using assms proof (safe intro!: measurable_comp[where b="prod_measure_space M1 M2"]) + interpret M1: algebra M1 by fact + interpret M2: algebra M2 by fact + assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2" + + show "f \<in> measurable M (prod_measure_space M1 M2)" unfolding prod_measure_space_def + proof (rule measurable_sigma) + show "prod_sets (sets M1) (sets M2) \<subseteq> Pow (space M1 \<times> space M2)" + unfolding prod_sets_def using M1.sets_into_space M2.sets_into_space by auto + show "f \<in> space M \<rightarrow> space M1 \<times> space M2" + using f s by (auto simp: mem_Times_iff measurable_def comp_def) + fix A assume "A \<in> prod_sets (sets M1) (sets M2)" + then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C" + unfolding prod_sets_def by auto + moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M" + using f `B \<in> sets M1` unfolding measurable_def by auto + moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M" + using s `C \<in> sets M2` unfolding measurable_def by auto + moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)" + unfolding `A = B \<times> C` by (auto simp: vimage_Times) + ultimately show "f -` A \<inter> space M \<in> sets M" by auto + qed +qed definition - "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))" + "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))" lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B" by (auto simp add: prod_sets_def) @@ -114,36 +512,25 @@ qed lemma (in finite_measure_space) finite_measure_space_finite_prod_measure: - assumes "finite_measure_space N \<nu>" + fixes N :: "('c, 'd) algebra_scheme" + assumes N: "finite_measure_space N \<nu>" shows "finite_measure_space (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)" unfolding finite_prod_measure_space[OF assms] -proof (rule finite_measure_spaceI) +proof (rule finite_measure_spaceI, simp_all) interpret N: finite_measure_space N \<nu> by fact - - let ?P = "\<lparr>space = space M \<times> space N, sets = Pow (space M \<times> space N)\<rparr>" - show "measure_space ?P (prod_measure M \<mu> N \<nu>)" - proof (rule sigma_algebra.finite_additivity_sufficient) - show "sigma_algebra ?P" by (rule sigma_algebra_Pow) - show "finite (space ?P)" using finite_space N.finite_space by auto - from finite_prod_measure_times[OF assms, of "{}" "{}"] - show "positive (prod_measure M \<mu> N \<nu>)" - unfolding positive_def by simp + show "finite (space M \<times> space N)" using finite_space N.finite_space by auto + show "prod_measure M \<mu> N \<nu> (space M \<times> space N) \<noteq> \<omega>" + using finite_prod_measure_times[OF N top N.top] by simp + show "prod_measure M \<mu> N \<nu> {} = 0" + using finite_prod_measure_times[OF N empty_sets N.empty_sets] by simp - show "additive ?P (prod_measure M \<mu> N \<nu>)" - unfolding additive_def - apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric] - intro!: positive_integral_cong) - apply (subst N.measure_additive[symmetric]) - by (auto simp: N.sets_eq_Pow sets_eq_Pow) - qed - show "finite (space ?P)" using finite_space N.finite_space by auto - show "sets ?P = Pow (space ?P)" by simp - - fix x assume "x \<in> space ?P" - with finite_prod_measure_times[OF assms, of "{fst x}" "{snd x}"] - finite_measure[of "{fst x}"] N.finite_measure[of "{snd x}"] - show "prod_measure M \<mu> N \<nu> {x} \<noteq> \<omega>" - by (auto simp add: sets_eq_Pow N.sets_eq_Pow elim!: SigmaE) + fix A B :: "('a * 'c) set" assume "A \<inter> B = {}" "A \<subseteq> space M \<times> space N" "B \<subseteq> space M \<times> space N" + then show "prod_measure M \<mu> N \<nu> (A \<union> B) = prod_measure M \<mu> N \<nu> A + prod_measure M \<mu> N \<nu> B" + apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric] + intro!: positive_integral_cong) + apply (subst N.measure_additive) + apply (auto intro!: arg_cong[where f=\<mu>] simp: N.sets_eq_Pow sets_eq_Pow) + done qed lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive: @@ -153,4 +540,18 @@ unfolding finite_prod_measure_space[OF N, symmetric] using finite_measure_space_finite_prod_measure[OF N] . -end \ No newline at end of file +lemma prod_measure_times_finite: + assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N" + shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}" +proof (cases a) + case (Pair b c) + hence a_eq: "{a} = {b} \<times> {c}" by simp + + interpret M: finite_measure_space M \<mu> by fact + interpret N: finite_measure_space N \<nu> by fact + + from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair + show ?thesis unfolding a_eq by simp +qed + +end
--- a/src/HOL/Probability/Radon_Nikodym.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Fri Sep 03 13:54:39 2010 +0200 @@ -2,201 +2,6 @@ imports Lebesgue_Integration begin -lemma (in measure_space) measure_finitely_subadditive: - assumes "finite I" "A ` I \<subseteq> sets M" - shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))" -using assms proof induct - case (insert i I) - then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN) - then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)" - using insert by (simp add: measure_subadditive) - also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))" - using insert by (auto intro!: add_left_mono) - finally show ?case . -qed simp - -lemma (in sigma_algebra) borel_measurable_restricted: - fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M" - shows "f \<in> borel_measurable (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<longleftrightarrow> - (\<lambda>x. f x * indicator A x) \<in> borel_measurable M" - (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M") -proof - - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) - have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R" - by (auto intro!: measurable_cong) - show ?thesis unfolding * - unfolding in_borel_measurable_borel_space - proof (simp, safe) - fix S :: "pinfreal set" assume "S \<in> sets borel_space" - "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M" - then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto - then have f: "?f -` S \<inter> A \<in> sets M" - using `A \<in> sets M` sets_into_space by fastsimp - show "?f -` S \<inter> space M \<in> sets M" - proof cases - assume "0 \<in> S" - then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)" - using `A \<in> sets M` sets_into_space by auto - then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff) - next - assume "0 \<notin> S" - then have "?f -` S \<inter> space M = ?f -` S \<inter> A" - using `A \<in> sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show ?thesis using f by auto - qed - next - fix S :: "pinfreal set" assume "S \<in> sets borel_space" - "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M" - then have f: "?f -` S \<inter> space M \<in> sets M" by auto - then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M" - using `A \<in> sets M` sets_into_space - apply (simp add: image_iff) - apply (rule bexI[OF _ f]) - by auto - qed -qed - -lemma (in sigma_algebra) simple_function_eq_borel_measurable: - fixes f :: "'a \<Rightarrow> pinfreal" - shows "simple_function f \<longleftrightarrow> - finite (f`space M) \<and> f \<in> borel_measurable M" - using simple_function_borel_measurable[of f] - borel_measurable_simple_function[of f] - by (fastsimp simp: simple_function_def) - -lemma (in measure_space) simple_function_restricted: - fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M" - shows "sigma_algebra.simple_function (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)" - (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f") -proof - - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) - have "finite (f`A) \<longleftrightarrow> finite (?f`space M)" - proof cases - assume "A = space M" - then have "f`A = ?f`space M" by (fastsimp simp: image_iff) - then show ?thesis by simp - next - assume "A \<noteq> space M" - then obtain x where x: "x \<in> space M" "x \<notin> A" - using sets_into_space `A \<in> sets M` by auto - have *: "?f`space M = f`A \<union> {0}" - proof (auto simp add: image_iff) - show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0" - using x by (auto intro!: bexI[of _ x]) - next - fix x assume "x \<in> A" - then show "\<exists>y\<in>space M. f x = f y * indicator A y" - using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x]) - next - fix x - assume "indicator A x \<noteq> (0::pinfreal)" - then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm) - moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y" - ultimately show "f x = 0" by auto - qed - then show ?thesis by auto - qed - then show ?thesis - unfolding simple_function_eq_borel_measurable - R.simple_function_eq_borel_measurable - unfolding borel_measurable_restricted[OF `A \<in> sets M`] - by auto -qed - -lemma (in measure_space) simple_integral_restricted: - assumes "A \<in> sets M" - assumes sf: "simple_function (\<lambda>x. f x * indicator A x)" - shows "measure_space.simple_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)" - (is "_ = simple_integral ?f") - unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]] - unfolding simple_integral_def -proof (simp, safe intro!: setsum_mono_zero_cong_left) - from sf show "finite (?f ` space M)" - unfolding simple_function_def by auto -next - fix x assume "x \<in> A" - then show "f x \<in> ?f ` space M" - using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x]) -next - fix x assume "x \<in> space M" "?f x \<notin> f`A" - then have "x \<notin> A" by (auto simp: image_iff) - then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp -next - fix x assume "x \<in> A" - then have "f x \<noteq> 0 \<Longrightarrow> - f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M" - using `A \<in> sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show "f x * \<mu> (f -` {f x} \<inter> A) = - f x * \<mu> (?f -` {f x} \<inter> space M)" - unfolding pinfreal_mult_cancel_left by auto -qed - -lemma (in measure_space) positive_integral_restricted: - assumes "A \<in> sets M" - shows "measure_space.positive_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)" - (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f") -proof - - have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`]) - then interpret R: measure_space ?R \<mu> . - have saR: "sigma_algebra ?R" by fact - have *: "R.positive_integral f = R.positive_integral ?f" - by (auto intro!: R.positive_integral_cong) - show ?thesis - unfolding * R.positive_integral_def positive_integral_def - unfolding simple_function_restricted[OF `A \<in> sets M`] - apply (simp add: SUPR_def) - apply (rule arg_cong[where f=Sup]) - proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`]) - fix g assume "simple_function (\<lambda>x. g x * indicator A x)" - "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x" - then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and> - simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x" - apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) - by (auto simp: indicator_def le_fun_def) - next - fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)" - "\<forall>x\<in>space M. \<omega> \<noteq> g x" - then have *: "(\<lambda>x. g x * indicator A x) = g" - "\<And>x. g x * indicator A x = g x" - "\<And>x. g x \<le> f x" - by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm) - from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and> - simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)" - using `A \<in> sets M`[THEN sets_into_space] - apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) - by (fastsimp simp: le_fun_def *) - qed -qed - -lemma (in sigma_algebra) borel_measurable_psuminf: - assumes "\<And>i. f i \<in> borel_measurable M" - shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M" - using assms unfolding psuminf_def - by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) - -lemma (in sigma_finite_measure) disjoint_sigma_finite: - "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> - (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A" -proof - - obtain A :: "nat \<Rightarrow> 'a set" where - range: "range A \<subseteq> sets M" and - space: "(\<Union>i. A i) = space M" and - measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" - using sigma_finite by auto - - note range' = range_disjointed_sets[OF range] range - - { fix i - have "\<mu> (disjointed A i) \<le> \<mu> (A i)" - using range' disjointed_subset[of A i] by (auto intro!: measure_mono) - then have "\<mu> (disjointed A i) \<noteq> \<omega>" - using measure[of i] by auto } - with disjoint_family_disjointed UN_disjointed_eq[of A] space range' - show ?thesis by (auto intro!: exI[of _ "disjointed A"]) -qed - lemma (in sigma_finite_measure) Ex_finite_integrable_function: shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)" proof - @@ -206,7 +11,6 @@ measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and disjoint: "disjoint_family A" using disjoint_sigma_finite by auto - let "?B i" = "2^Suc i * \<mu> (A i)" have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)" proof @@ -225,20 +29,22 @@ qed from choice[OF this] obtain n where n: "\<And>i. 0 < n i" "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto - let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x" show ?thesis proof (safe intro!: bexI[of _ ?h] del: notI) - have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))" - apply (subst positive_integral_psuminf) - using range apply (fastsimp intro!: borel_measurable_pinfreal_times borel_measurable_const borel_measurable_indicator) - apply (subst positive_integral_cmult_indicator) - using range by auto + have "\<And>i. A i \<in> sets M" + using range by fastsimp+ + then have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))" + by (simp add: positive_integral_psuminf positive_integral_cmult_indicator) also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))" proof (rule psuminf_le) fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)" using measure[of N] n[of N] - by (cases "n N") (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff mult_le_0_iff mult_less_0_iff power_less_zero_eq power_le_zero_eq inverse_eq_divide less_divide_eq power_divide split: split_if_asm) + by (cases "n N") + (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff + mult_le_0_iff mult_less_0_iff power_less_zero_eq + power_le_zero_eq inverse_eq_divide less_divide_eq + power_divide split: split_if_asm) qed also have "\<dots> = Real 1" by (rule suminf_imp_psuminf, rule power_half_series, auto) @@ -251,13 +57,37 @@ then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto next show "?h \<in> borel_measurable M" using range - by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times borel_measurable_indicator) + by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times) qed qed definition (in measure_space) "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))" +lemma (in finite_measure_space) absolutely_continuousI: + assumes "finite_measure_space M \<nu>" + assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0" + shows "absolutely_continuous \<nu>" +proof (unfold absolutely_continuous_def sets_eq_Pow, safe) + fix N assume "\<mu> N = 0" "N \<subseteq> space M" + interpret v: finite_measure_space M \<nu> by fact + have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp + also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})" + proof (rule v.measure_finitely_additive''[symmetric]) + show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset) + show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto + fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto + qed + also have "\<dots> = 0" + proof (safe intro!: setsum_0') + fix x assume "x \<in> N" + hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono) + hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp + thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto + qed + finally show "\<nu> N = 0" . +qed + lemma (in finite_measure) Radon_Nikodym_aux_epsilon: fixes e :: real assumes "0 < e" assumes "finite_measure M s" @@ -370,7 +200,7 @@ interpret M': finite_measure M s by fact - let "?r S" = "M\<lparr> space := S, sets := (\<lambda>C. S \<inter> C)`sets M\<rparr>" + let "?r S" = "restricted_space S" { fix S n assume S: "S \<in> sets M" @@ -838,7 +668,7 @@ = (f x * indicator (Q i) x) * indicator A x" unfolding indicator_def by auto - have fm: "finite_measure (M\<lparr>space := Q i, sets := op \<inter> (Q i) ` sets M\<rparr>) \<mu>" + have fm: "finite_measure (restricted_space (Q i)) \<mu>" (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]]) then interpret R: finite_measure ?R . have fmv: "finite_measure ?R \<nu>" @@ -935,47 +765,6 @@ qed qed -lemma (in measure_space) positive_integral_translated_density: - assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" - shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g = - positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") -proof - - from measure_space_density[OF assms(1)] - interpret T: measure_space M ?T . - - from borel_measurable_implies_simple_function_sequence[OF assms(2)] - obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast - note G_borel = borel_measurable_simple_function[OF this(1)] - - from T.positive_integral_isoton[OF `G \<up> g` G_borel] - have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" . - - { fix i - have [simp]: "finite (G i ` space M)" - using G(1) unfolding simple_function_def by auto - have "T.positive_integral (G i) = T.simple_integral (G i)" - using G T.positive_integral_eq_simple_integral by simp - also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))" - apply (simp add: T.simple_integral_def) - apply (subst positive_integral_cmult[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) - apply (subst positive_integral_setsum[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) - by (simp add: setsum_right_distrib field_simps) - also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)" - by (auto intro!: positive_integral_cong - simp: indicator_def if_distrib setsum_cases) - finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . } - with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp - - from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)" - unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) - then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)" - using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times) - with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)" - unfolding isoton_def by simp -qed - lemma (in sigma_finite_measure) Radon_Nikodym: assumes "measure_space M \<nu>" assumes "absolutely_continuous \<nu>"
--- a/src/HOL/Probability/Sigma_Algebra.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Fri Sep 03 13:54:39 2010 +0200 @@ -6,7 +6,7 @@ header {* Sigma Algebras *} -theory Sigma_Algebra imports Main Countable FuncSet begin +theory Sigma_Algebra imports Main Countable FuncSet Indicator_Function begin text {* Sigma algebras are an elementary concept in measure theory. To measure --- that is to integrate --- functions, we first have @@ -95,10 +95,13 @@ lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x" by (metis Int_absorb2 sets_into_space) +section {* Restricted algebras *} + +abbreviation (in algebra) + "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M \<rparr>" + lemma (in algebra) restricted_algebra: - assumes "S \<in> sets M" - shows "algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)" - (is "algebra ?r") + assumes "A \<in> sets M" shows "algebra (restricted_space A)" using assms by unfold_locales auto subsection {* Sigma Algebras *} @@ -107,6 +110,13 @@ assumes countable_nat_UN [intro]: "!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M" +lemma sigma_algebra_cong: + fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme" + assumes *: "sigma_algebra M" + and cong: "space M = space M'" "sets M = sets M'" + shows "sigma_algebra M'" +using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong . + lemma countable_UN_eq: fixes A :: "'i::countable \<Rightarrow> 'a set" shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow> @@ -320,15 +330,14 @@ lemma (in sigma_algebra) restricted_sigma_algebra: assumes "S \<in> sets M" - shows "sigma_algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)" - (is "sigma_algebra ?r") + shows "sigma_algebra (restricted_space S)" unfolding sigma_algebra_def sigma_algebra_axioms_def proof safe - show "algebra ?r" using restricted_algebra[OF assms] . + show "algebra (restricted_space S)" using restricted_algebra[OF assms] . next - fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets ?r" + fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (restricted_space S)" from restriction_in_sets[OF assms this[simplified]] - show "(\<Union>i. A i) \<in> sets ?r" by simp + show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp qed section {* Measurable functions *} @@ -560,6 +569,19 @@ (metis insert_absorb insert_subset le_SucE le_antisym not_leE) qed +lemma setsum_indicator_disjoint_family: + fixes f :: "'d \<Rightarrow> 'e::semiring_1" + assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P" + shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j" +proof - + have "P \<inter> {i. x \<in> A i} = {j}" + using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def + by auto + thus ?thesis + unfolding indicator_def + by (simp add: if_distrib setsum_cases[OF `finite P`]) +qed + definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set " where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)" @@ -626,6 +648,67 @@ thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq) qed +subsection {* Sigma algebra generated by function preimages *} + +definition (in sigma_algebra) + "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M \<rparr>" + +lemma (in sigma_algebra) in_vimage_algebra[simp]: + "A \<in> sets (vimage_algebra S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)" + by (simp add: vimage_algebra_def image_iff) + +lemma (in sigma_algebra) space_vimage_algebra[simp]: + "space (vimage_algebra S f) = S" + by (simp add: vimage_algebra_def) + +lemma (in sigma_algebra) sigma_algebra_vimage: + fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M" + shows "sigma_algebra (vimage_algebra S f)" +proof + fix A assume "A \<in> sets (vimage_algebra S f)" + then guess B unfolding in_vimage_algebra .. + then show "space (vimage_algebra S f) - A \<in> sets (vimage_algebra S f)" + using sets_into_space assms + by (auto intro!: bexI[of _ "space M - B"]) +next + fix A assume "A \<in> sets (vimage_algebra S f)" + then guess A' unfolding in_vimage_algebra .. note A' = this + fix B assume "B \<in> sets (vimage_algebra S f)" + then guess B' unfolding in_vimage_algebra .. note B' = this + then show "A \<union> B \<in> sets (vimage_algebra S f)" + using sets_into_space assms A' B' + by (auto intro!: bexI[of _ "A' \<union> B'"]) +next + fix A::"nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets (vimage_algebra S f)" + then have "\<forall>i. \<exists>B. A i = f -` B \<inter> S \<and> B \<in> sets M" + by (simp add: subset_eq) blast + from this[THEN choice] obtain B + where B: "\<And>i. A i = f -` B i \<inter> S" "range B \<subseteq> sets M" by auto + show "(\<Union>i. A i) \<in> sets (vimage_algebra S f)" + using B by (auto intro!: bexI[of _ "\<Union>i. B i"]) +qed auto + +lemma (in sigma_algebra) measurable_vimage_algebra: + fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M" + shows "f \<in> measurable (vimage_algebra S f) M" + unfolding measurable_def using assms by force + +section {* Conditional space *} + +definition (in algebra) + "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M \<rparr>" + +definition (in algebra) + "conditional_space X A = algebra.image_space (restricted_space A) X" + +lemma (in algebra) space_conditional_space: + assumes "A \<in> sets M" shows "space (conditional_space X A) = X`A" +proof - + interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra) + show ?thesis unfolding conditional_space_def r.image_space_def + by simp +qed + subsection {* A Two-Element Series *} definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Fri Sep 03 13:54:39 2010 +0200 @@ -26,14 +26,13 @@ let ?measure = "\<lambda>s::'a set. real (card s) / real (card S)" show "finite_measure_space M \<mu>" - proof (rule finite_Pow_additivity_sufficient, simp_all) - show "positive \<mu>" by (simp add: positive_def) - - show "additive M \<mu>" - by (simp add: additive_def inverse_eq_divide field_simps Real_real + proof (rule finite_measure_spaceI) + fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M" + then show "\<mu> (A \<union> B) = \<mu> A + \<mu> B" + by (simp add: inverse_eq_divide field_simps Real_real divide_le_0_iff zero_le_divide_iff card_Un_disjoint finite_subset[OF _ finite]) - qed + qed auto qed simp_all lemma set_of_list_extend:
--- a/src/HOL/Refute.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Refute.thy Fri Sep 03 13:54:39 2010 +0200 @@ -9,9 +9,7 @@ theory Refute imports Hilbert_Choice List -uses - "Tools/refute.ML" - "Tools/refute_isar.ML" +uses "Tools/refute.ML" begin setup Refute.setup @@ -92,16 +90,14 @@ (* ------------------------------------------------------------------------- *) (* FILES *) (* *) -(* HOL/Tools/prop_logic.ML Propositional logic *) -(* HOL/Tools/sat_solver.ML SAT solvers *) -(* HOL/Tools/refute.ML Translation HOL -> propositional logic and *) -(* Boolean assignment -> HOL model *) -(* HOL/Tools/refute_isar.ML Adds 'refute'/'refute_params' to Isabelle's *) -(* syntax *) -(* HOL/Refute.thy This file: loads the ML files, basic setup, *) -(* documentation *) -(* HOL/SAT.thy Sets default parameters *) -(* HOL/ex/RefuteExamples.thy Examples *) +(* HOL/Tools/prop_logic.ML Propositional logic *) +(* HOL/Tools/sat_solver.ML SAT solvers *) +(* HOL/Tools/refute.ML Translation HOL -> propositional logic and *) +(* Boolean assignment -> HOL model *) +(* HOL/Refute.thy This file: loads the ML files, basic setup, *) +(* documentation *) +(* HOL/SAT.thy Sets default parameters *) +(* HOL/ex/Refute_Examples.thy Examples *) (* ------------------------------------------------------------------------- *) \end{verbatim} *}
--- a/src/HOL/SetInterval.thy Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/SetInterval.thy Fri Sep 03 13:54:39 2010 +0200 @@ -304,6 +304,17 @@ lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) +text {* The following proof is convinient in induction proofs where +new elements get indices at the beginning. So it is used to transform +@{term "{..<Suc n}"} to @{term "0::nat"} and @{term "{..< n}"}. *} + +lemma lessThan_Suc_eq_insert_0: "{..<Suc n} = insert 0 (Suc ` {..<n})" +proof safe + fix x assume "x < Suc n" "x \<notin> Suc ` {..<n}" + then have "x \<noteq> Suc (x - 1)" by auto + with `x < Suc n` show "x = 0" by auto +qed + lemma lessThan_Suc_atMost: "lessThan (Suc k) = atMost k" by (simp add: lessThan_def atMost_def less_Suc_eq_le)
--- a/src/HOL/Tools/refute.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Tools/refute.ML Fri Sep 03 13:54:39 2010 +0200 @@ -25,34 +25,34 @@ exception MAXVARS_EXCEEDED - val add_interpreter : string -> (theory -> model -> arguments -> term -> + val add_interpreter : string -> (Proof.context -> model -> arguments -> term -> (interpretation * model * arguments) option) -> theory -> theory - val add_printer : string -> (theory -> model -> typ -> + val add_printer : string -> (Proof.context -> model -> typ -> interpretation -> (int -> bool) -> term option) -> theory -> theory - val interpret : theory -> model -> arguments -> term -> + val interpret : Proof.context -> model -> arguments -> term -> (interpretation * model * arguments) - val print : theory -> model -> typ -> interpretation -> (int -> bool) -> term - val print_model : theory -> model -> (int -> bool) -> string + val print : Proof.context -> model -> typ -> interpretation -> (int -> bool) -> term + val print_model : Proof.context -> model -> (int -> bool) -> string (* ------------------------------------------------------------------------- *) (* Interface *) (* ------------------------------------------------------------------------- *) val set_default_param : (string * string) -> theory -> theory - val get_default_param : theory -> string -> string option - val get_default_params : theory -> (string * string) list - val actual_params : theory -> (string * string) list -> params + val get_default_param : Proof.context -> string -> string option + val get_default_params : Proof.context -> (string * string) list + val actual_params : Proof.context -> (string * string) list -> params - val find_model : theory -> params -> term list -> term -> bool -> unit + val find_model : Proof.context -> params -> term list -> term -> bool -> unit (* tries to find a model for a formula: *) val satisfy_term : - theory -> (string * string) list -> term list -> term -> unit + Proof.context -> (string * string) list -> term list -> term -> unit (* tries to find a model that refutes a formula: *) val refute_term : - theory -> (string * string) list -> term list -> term -> unit + Proof.context -> (string * string) list -> term list -> term -> unit val refute_goal : Proof.context -> (string * string) list -> thm -> int -> unit @@ -72,22 +72,23 @@ val is_const_of_class: theory -> string * typ -> bool val string_of_typ : typ -> string val typ_of_dtyp : Datatype.descr -> (Datatype.dtyp * typ) list -> Datatype.dtyp -> typ -end; (* signature REFUTE *) +end; structure Refute : REFUTE = struct - open PropLogic; +open PropLogic; - (* We use 'REFUTE' only for internal error conditions that should *) - (* never occur in the first place (i.e. errors caused by bugs in our *) - (* code). Otherwise (e.g. to indicate invalid input data) we use *) - (* 'error'. *) - exception REFUTE of string * string; (* ("in function", "cause") *) +(* We use 'REFUTE' only for internal error conditions that should *) +(* never occur in the first place (i.e. errors caused by bugs in our *) +(* code). Otherwise (e.g. to indicate invalid input data) we use *) +(* 'error'. *) +exception REFUTE of string * string; (* ("in function", "cause") *) - (* should be raised by an interpreter when more variables would be *) - (* required than allowed by 'maxvars' *) - exception MAXVARS_EXCEEDED; +(* should be raised by an interpreter when more variables would be *) +(* required than allowed by 'maxvars' *) +exception MAXVARS_EXCEEDED; + (* ------------------------------------------------------------------------- *) (* TREES *) @@ -98,20 +99,20 @@ (* of (lists of ...) elements *) (* ------------------------------------------------------------------------- *) - datatype 'a tree = - Leaf of 'a - | Node of ('a tree) list; +datatype 'a tree = + Leaf of 'a + | Node of ('a tree) list; - (* ('a -> 'b) -> 'a tree -> 'b tree *) +(* ('a -> 'b) -> 'a tree -> 'b tree *) - fun tree_map f tr = - case tr of - Leaf x => Leaf (f x) - | Node xs => Node (map (tree_map f) xs); +fun tree_map f tr = + case tr of + Leaf x => Leaf (f x) + | Node xs => Node (map (tree_map f) xs); - (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *) +(* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *) - fun tree_foldl f = +fun tree_foldl f = let fun itl (e, Leaf x) = f(e,x) | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs) @@ -119,16 +120,16 @@ itl end; - (* 'a tree * 'b tree -> ('a * 'b) tree *) +(* 'a tree * 'b tree -> ('a * 'b) tree *) - fun tree_pair (t1, t2) = - case t1 of - Leaf x => +fun tree_pair (t1, t2) = + case t1 of + Leaf x => (case t2 of Leaf y => Leaf (x,y) | Node _ => raise REFUTE ("tree_pair", "trees are of different height (second tree is higher)")) - | Node xs => + | Node xs => (case t2 of (* '~~' will raise an exception if the number of branches in *) (* both trees is different at the current node *) @@ -160,68 +161,70 @@ (* "unknown"). *) (* ------------------------------------------------------------------------- *) - type params = - { - sizes : (string * int) list, - minsize : int, - maxsize : int, - maxvars : int, - maxtime : int, - satsolver: string, - no_assms : bool, - expect : string - }; +type params = + { + sizes : (string * int) list, + minsize : int, + maxsize : int, + maxvars : int, + maxtime : int, + satsolver: string, + no_assms : bool, + expect : string + }; (* ------------------------------------------------------------------------- *) (* interpretation: a term's interpretation is given by a variable of type *) (* 'interpretation' *) (* ------------------------------------------------------------------------- *) - type interpretation = - prop_formula list tree; +type interpretation = + prop_formula list tree; (* ------------------------------------------------------------------------- *) (* model: a model specifies the size of types and the interpretation of *) (* terms *) (* ------------------------------------------------------------------------- *) - type model = - (typ * int) list * (term * interpretation) list; +type model = + (typ * int) list * (term * interpretation) list; (* ------------------------------------------------------------------------- *) (* arguments: additional arguments required during interpretation of terms *) (* ------------------------------------------------------------------------- *) - type arguments = - { - (* just passed unchanged from 'params': *) - maxvars : int, - (* whether to use 'make_equality' or 'make_def_equality': *) - def_eq : bool, - (* the following may change during the translation: *) - next_idx : int, - bounds : interpretation list, - wellformed: prop_formula - }; +type arguments = + { + (* just passed unchanged from 'params': *) + maxvars : int, + (* whether to use 'make_equality' or 'make_def_equality': *) + def_eq : bool, + (* the following may change during the translation: *) + next_idx : int, + bounds : interpretation list, + wellformed: prop_formula + }; - structure RefuteData = Theory_Data - ( - type T = - {interpreters: (string * (theory -> model -> arguments -> term -> - (interpretation * model * arguments) option)) list, - printers: (string * (theory -> model -> typ -> interpretation -> - (int -> bool) -> term option)) list, - parameters: string Symtab.table}; - val empty = {interpreters = [], printers = [], parameters = Symtab.empty}; - val extend = I; - fun merge - ({interpreters = in1, printers = pr1, parameters = pa1}, - {interpreters = in2, printers = pr2, parameters = pa2}) : T = - {interpreters = AList.merge (op =) (K true) (in1, in2), - printers = AList.merge (op =) (K true) (pr1, pr2), - parameters = Symtab.merge (op=) (pa1, pa2)}; - ); +structure Data = Theory_Data +( + type T = + {interpreters: (string * (Proof.context -> model -> arguments -> term -> + (interpretation * model * arguments) option)) list, + printers: (string * (Proof.context -> model -> typ -> interpretation -> + (int -> bool) -> term option)) list, + parameters: string Symtab.table}; + val empty = {interpreters = [], printers = [], parameters = Symtab.empty}; + val extend = I; + fun merge + ({interpreters = in1, printers = pr1, parameters = pa1}, + {interpreters = in2, printers = pr2, parameters = pa2}) : T = + {interpreters = AList.merge (op =) (K true) (in1, in2), + printers = AList.merge (op =) (K true) (pr1, pr2), + parameters = Symtab.merge (op=) (pa1, pa2)}; +); + +val get_data = Data.get o ProofContext.theory_of; (* ------------------------------------------------------------------------- *) @@ -230,30 +233,24 @@ (* track of the interpretation of subterms *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) *) - - fun interpret thy model args t = - case get_first (fn (_, f) => f thy model args t) - (#interpreters (RefuteData.get thy)) of - NONE => raise REFUTE ("interpret", - "no interpreter for term " ^ quote (Syntax.string_of_term_global thy t)) - | SOME x => x; +fun interpret ctxt model args t = + case get_first (fn (_, f) => f ctxt model args t) + (#interpreters (get_data ctxt)) of + NONE => raise REFUTE ("interpret", + "no interpreter for term " ^ quote (Syntax.string_of_term ctxt t)) + | SOME x => x; (* ------------------------------------------------------------------------- *) (* print: converts the interpretation 'intr', which must denote a term of *) (* type 'T', into a term using a suitable printer *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> - Term.term *) - - fun print thy model T intr assignment = - case get_first (fn (_, f) => f thy model T intr assignment) - (#printers (RefuteData.get thy)) of - NONE => raise REFUTE ("print", - "no printer for type " ^ quote (Syntax.string_of_typ_global thy T)) - | SOME x => x; +fun print ctxt model T intr assignment = + case get_first (fn (_, f) => f ctxt model T intr assignment) + (#printers (get_data ctxt)) of + NONE => raise REFUTE ("print", + "no printer for type " ^ quote (Syntax.string_of_typ ctxt T)) + | SOME x => x; (* ------------------------------------------------------------------------- *) (* print_model: turns the model into a string, using a fixed interpretation *) @@ -261,9 +258,7 @@ (* printers *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> (int -> bool) -> string *) - - fun print_model thy model assignment = +fun print_model ctxt model assignment = let val (typs, terms) = model val typs_msg = @@ -271,10 +266,10 @@ "empty universe (no type variables in term)\n" else "Size of types: " ^ commas (map (fn (T, i) => - Syntax.string_of_typ_global thy T ^ ": " ^ string_of_int i) typs) ^ "\n" + Syntax.string_of_typ ctxt T ^ ": " ^ string_of_int i) typs) ^ "\n" val show_consts_msg = - if not (!show_consts) andalso Library.exists (is_Const o fst) terms then - "set \"show_consts\" to show the interpretation of constants\n" + if not (Config.get ctxt show_consts) andalso Library.exists (is_Const o fst) terms then + "enable \"show_consts\" to show the interpretation of constants\n" else "" val terms_msg = @@ -283,10 +278,10 @@ else cat_lines (map_filter (fn (t, intr) => (* print constants only if 'show_consts' is true *) - if (!show_consts) orelse not (is_Const t) then - SOME (Syntax.string_of_term_global thy t ^ ": " ^ - Syntax.string_of_term_global thy - (print thy model (Term.type_of t) intr assignment)) + if Config.get ctxt show_consts orelse not (is_Const t) then + SOME (Syntax.string_of_term ctxt t ^ ": " ^ + Syntax.string_of_term ctxt + (print ctxt model (Term.type_of t) intr assignment)) else NONE) terms) ^ "\n" in @@ -298,71 +293,49 @@ (* PARAMETER MANAGEMENT *) (* ------------------------------------------------------------------------- *) - (* string -> (theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option) -> theory -> theory *) - - fun add_interpreter name f thy = - let - val {interpreters, printers, parameters} = RefuteData.get thy - in - case AList.lookup (op =) interpreters name of - NONE => RefuteData.put {interpreters = (name, f) :: interpreters, - printers = printers, parameters = parameters} thy - | SOME _ => error ("Interpreter " ^ name ^ " already declared") - end; +fun add_interpreter name f = Data.map (fn {interpreters, printers, parameters} => + case AList.lookup (op =) interpreters name of + NONE => {interpreters = (name, f) :: interpreters, + printers = printers, parameters = parameters} + | SOME _ => error ("Interpreter " ^ name ^ " already declared")); - (* string -> (theory -> model -> Term.typ -> interpretation -> - (int -> bool) -> Term.term option) -> theory -> theory *) - - fun add_printer name f thy = - let - val {interpreters, printers, parameters} = RefuteData.get thy - in - case AList.lookup (op =) printers name of - NONE => RefuteData.put {interpreters = interpreters, - printers = (name, f) :: printers, parameters = parameters} thy - | SOME _ => error ("Printer " ^ name ^ " already declared") - end; +fun add_printer name f = Data.map (fn {interpreters, printers, parameters} => + case AList.lookup (op =) printers name of + NONE => {interpreters = interpreters, + printers = (name, f) :: printers, parameters = parameters} + | SOME _ => error ("Printer " ^ name ^ " already declared")); (* ------------------------------------------------------------------------- *) -(* set_default_param: stores the '(name, value)' pair in RefuteData's *) +(* set_default_param: stores the '(name, value)' pair in Data's *) (* parameter table *) (* ------------------------------------------------------------------------- *) - (* (string * string) -> theory -> theory *) - - fun set_default_param (name, value) = RefuteData.map - (fn {interpreters, printers, parameters} => - {interpreters = interpreters, printers = printers, - parameters = Symtab.update (name, value) parameters}); +fun set_default_param (name, value) = Data.map + (fn {interpreters, printers, parameters} => + {interpreters = interpreters, printers = printers, + parameters = Symtab.update (name, value) parameters}); (* ------------------------------------------------------------------------- *) (* get_default_param: retrieves the value associated with 'name' from *) -(* RefuteData's parameter table *) +(* Data's parameter table *) (* ------------------------------------------------------------------------- *) - (* theory -> string -> string option *) - - val get_default_param = Symtab.lookup o #parameters o RefuteData.get; +val get_default_param = Symtab.lookup o #parameters o get_data; (* ------------------------------------------------------------------------- *) (* get_default_params: returns a list of all '(name, value)' pairs that are *) -(* stored in RefuteData's parameter table *) +(* stored in Data's parameter table *) (* ------------------------------------------------------------------------- *) - (* theory -> (string * string) list *) - - val get_default_params = Symtab.dest o #parameters o RefuteData.get; +val get_default_params = Symtab.dest o #parameters o get_data; (* ------------------------------------------------------------------------- *) (* actual_params: takes a (possibly empty) list 'params' of parameters that *) -(* override the default parameters currently specified in 'thy', and *) +(* override the default parameters currently specified, and *) (* returns a record that can be passed to 'find_model'. *) (* ------------------------------------------------------------------------- *) - (* theory -> (string * string) list -> params *) - - fun actual_params thy override = +fun actual_params ctxt override = let (* (string * string) list * string -> bool *) fun read_bool (parms, name) = @@ -370,32 +343,33 @@ SOME "true" => true | SOME "false" => false | SOME s => error ("parameter " ^ quote name ^ - " (value is " ^ quote s ^ ") must be \"true\" or \"false\"") + " (value is " ^ quote s ^ ") must be \"true\" or \"false\"") | NONE => error ("parameter " ^ quote name ^ " must be assigned a value") (* (string * string) list * string -> int *) fun read_int (parms, name) = case AList.lookup (op =) parms name of - SOME s => (case Int.fromString s of - SOME i => i - | NONE => error ("parameter " ^ quote name ^ - " (value is " ^ quote s ^ ") must be an integer value")) - | NONE => error ("parameter " ^ quote name ^ + SOME s => + (case Int.fromString s of + SOME i => i + | NONE => error ("parameter " ^ quote name ^ + " (value is " ^ quote s ^ ") must be an integer value")) + | NONE => error ("parameter " ^ quote name ^ " must be assigned a value") (* (string * string) list * string -> string *) fun read_string (parms, name) = case AList.lookup (op =) parms name of SOME s => s - | NONE => error ("parameter " ^ quote name ^ + | NONE => error ("parameter " ^ quote name ^ " must be assigned a value") (* 'override' first, defaults last: *) (* (string * string) list *) - val allparams = override @ (get_default_params thy) + val allparams = override @ get_default_params ctxt (* int *) - val minsize = read_int (allparams, "minsize") - val maxsize = read_int (allparams, "maxsize") - val maxvars = read_int (allparams, "maxvars") - val maxtime = read_int (allparams, "maxtime") + val minsize = read_int (allparams, "minsize") + val maxsize = read_int (allparams, "maxsize") + val maxvars = read_int (allparams, "maxvars") + val maxtime = read_int (allparams, "maxtime") (* string *) val satsolver = read_string (allparams, "satsolver") val no_assms = read_bool (allparams, "no_assms") @@ -405,7 +379,7 @@ (* TODO: it is currently not possible to specify a size for a type *) (* whose name is one of the other parameters (e.g. 'maxvars') *) (* (string * int) list *) - val sizes = map_filter + val sizes = map_filter (fn (name, value) => Option.map (pair name) (Int.fromString value)) (filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" @@ -420,25 +394,25 @@ (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *) (* ------------------------------------------------------------------------- *) - fun typ_of_dtyp descr typ_assoc (Datatype_Aux.DtTFree a) = - (* replace a 'DtTFree' variable by the associated type *) - the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a)) - | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, ds)) = - Type (s, map (typ_of_dtyp descr typ_assoc) ds) - | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) = - let - val (s, ds, _) = the (AList.lookup (op =) descr i) - in +fun typ_of_dtyp descr typ_assoc (Datatype_Aux.DtTFree a) = + (* replace a 'DtTFree' variable by the associated type *) + the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a)) + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, ds)) = Type (s, map (typ_of_dtyp descr typ_assoc) ds) - end; + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) = + let + val (s, ds, _) = the (AList.lookup (op =) descr i) + in + Type (s, map (typ_of_dtyp descr typ_assoc) ds) + end; (* ------------------------------------------------------------------------- *) (* close_form: universal closure over schematic variables in 't' *) (* ------------------------------------------------------------------------- *) - (* Term.term -> Term.term *) +(* Term.term -> Term.term *) - fun close_form t = +fun close_form t = let (* (Term.indexname * Term.typ) list *) val vars = sort_wrt (fst o fst) (map dest_Var (OldTerm.term_vars t)) @@ -455,9 +429,7 @@ (* denotes membership to an axiomatic type class *) (* ------------------------------------------------------------------------- *) - (* theory -> string * Term.typ -> bool *) - - fun is_const_of_class thy (s, T) = +fun is_const_of_class thy (s, T) = let val class_const_names = map Logic.const_of_class (Sign.all_classes thy) in @@ -471,28 +443,22 @@ (* of an inductive datatype in 'thy' *) (* ------------------------------------------------------------------------- *) - (* theory -> string * Term.typ -> bool *) - - fun is_IDT_constructor thy (s, T) = - (case body_type T of - Type (s', _) => +fun is_IDT_constructor thy (s, T) = + (case body_type T of + Type (s', _) => (case Datatype.get_constrs thy s' of SOME constrs => - List.exists (fn (cname, cty) => - cname = s andalso Sign.typ_instance thy (T, cty)) constrs - | NONE => - false) - | _ => - false); + List.exists (fn (cname, cty) => + cname = s andalso Sign.typ_instance thy (T, cty)) constrs + | NONE => false) + | _ => false); (* ------------------------------------------------------------------------- *) (* is_IDT_recursor: returns 'true' iff 'Const (s, T)' is the recursion *) (* operator of an inductive datatype in 'thy' *) (* ------------------------------------------------------------------------- *) - (* theory -> string * Term.typ -> bool *) - - fun is_IDT_recursor thy (s, T) = +fun is_IDT_recursor thy (s, T) = let val rec_names = Symtab.fold (append o #rec_names o snd) (Datatype.get_all thy) [] @@ -506,12 +472,12 @@ (* norm_rhs: maps f ?t1 ... ?tn == rhs to %t1...tn. rhs *) (* ------------------------------------------------------------------------- *) - fun norm_rhs eqn = +fun norm_rhs eqn = let fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t)) - | lambda v t = raise TERM ("lambda", [v, t]) + | lambda v t = raise TERM ("lambda", [v, t]) val (lhs, rhs) = Logic.dest_equals eqn - val (_, args) = Term.strip_comb lhs + val (_, args) = Term.strip_comb lhs in fold lambda (rev args) rhs end @@ -520,31 +486,29 @@ (* get_def: looks up the definition of a constant *) (* ------------------------------------------------------------------------- *) - (* theory -> string * Term.typ -> (string * Term.term) option *) - - fun get_def thy (s, T) = +fun get_def thy (s, T) = let (* (string * Term.term) list -> (string * Term.term) option *) fun get_def_ax [] = NONE | get_def_ax ((axname, ax) :: axioms) = - (let - val (lhs, _) = Logic.dest_equals ax (* equations only *) - val c = Term.head_of lhs - val (s', T') = Term.dest_Const c - in - if s=s' then - let - val typeSubs = Sign.typ_match thy (T', T) Vartab.empty - val ax' = monomorphic_term typeSubs ax - val rhs = norm_rhs ax' + (let + val (lhs, _) = Logic.dest_equals ax (* equations only *) + val c = Term.head_of lhs + val (s', T') = Term.dest_Const c in - SOME (axname, rhs) - end - else - get_def_ax axioms - end handle ERROR _ => get_def_ax axioms - | TERM _ => get_def_ax axioms - | Type.TYPE_MATCH => get_def_ax axioms) + if s=s' then + let + val typeSubs = Sign.typ_match thy (T', T) Vartab.empty + val ax' = monomorphic_term typeSubs ax + val rhs = norm_rhs ax' + in + SOME (axname, rhs) + end + else + get_def_ax axioms + end handle ERROR _ => get_def_ax axioms + | TERM _ => get_def_ax axioms + | Type.TYPE_MATCH => get_def_ax axioms) in get_def_ax (Theory.all_axioms_of thy) end; @@ -553,43 +517,40 @@ (* get_typedef: looks up the definition of a type, as created by "typedef" *) (* ------------------------------------------------------------------------- *) - (* theory -> Term.typ -> (string * Term.term) option *) - - fun get_typedef thy T = +fun get_typedef thy T = let (* (string * Term.term) list -> (string * Term.term) option *) fun get_typedef_ax [] = NONE | get_typedef_ax ((axname, ax) :: axioms) = - (let - (* Term.term -> Term.typ option *) - fun type_of_type_definition (Const (s', T')) = - if s'= @{const_name type_definition} then - SOME T' - else - NONE - | type_of_type_definition (Free _) = NONE - | type_of_type_definition (Var _) = NONE - | type_of_type_definition (Bound _) = NONE - | type_of_type_definition (Abs (_, _, body)) = - type_of_type_definition body - | type_of_type_definition (t1 $ t2) = - (case type_of_type_definition t1 of - SOME x => SOME x - | NONE => type_of_type_definition t2) - in - case type_of_type_definition ax of - SOME T' => - let - val T'' = (domain_type o domain_type) T' - val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty + (let + (* Term.term -> Term.typ option *) + fun type_of_type_definition (Const (s', T')) = + if s'= @{const_name type_definition} then + SOME T' + else + NONE + | type_of_type_definition (Free _) = NONE + | type_of_type_definition (Var _) = NONE + | type_of_type_definition (Bound _) = NONE + | type_of_type_definition (Abs (_, _, body)) = + type_of_type_definition body + | type_of_type_definition (t1 $ t2) = + (case type_of_type_definition t1 of + SOME x => SOME x + | NONE => type_of_type_definition t2) in - SOME (axname, monomorphic_term typeSubs ax) - end - | NONE => - get_typedef_ax axioms - end handle ERROR _ => get_typedef_ax axioms - | TERM _ => get_typedef_ax axioms - | Type.TYPE_MATCH => get_typedef_ax axioms) + case type_of_type_definition ax of + SOME T' => + let + val T'' = domain_type (domain_type T') + val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty + in + SOME (axname, monomorphic_term typeSubs ax) + end + | NONE => get_typedef_ax axioms + end handle ERROR _ => get_typedef_ax axioms + | TERM _ => get_typedef_ax axioms + | Type.TYPE_MATCH => get_typedef_ax axioms) in get_typedef_ax (Theory.all_axioms_of thy) end; @@ -599,9 +560,7 @@ (* created by the "axclass" command *) (* ------------------------------------------------------------------------- *) - (* theory -> string -> (string * Term.term) option *) - - fun get_classdef thy class = +fun get_classdef thy class = let val axname = class ^ "_class_def" in @@ -617,15 +576,13 @@ (* that definition does not need to be unfolded *) (* ------------------------------------------------------------------------- *) - (* theory -> Term.term -> Term.term *) +(* Note: we could intertwine unfolding of constants and beta-(eta-) *) +(* normalization; this would save some unfolding for terms where *) +(* constants are eliminated by beta-reduction (e.g. 'K c1 c2'). On *) +(* the other hand, this would cause additional work for terms where *) +(* constants are duplicated by beta-reduction (e.g. 'S c1 c2 c3'). *) - (* Note: we could intertwine unfolding of constants and beta-(eta-) *) - (* normalization; this would save some unfolding for terms where *) - (* constants are eliminated by beta-reduction (e.g. 'K c1 c2'). On *) - (* the other hand, this would cause additional work for terms where *) - (* constants are duplicated by beta-reduction (e.g. 'S c1 c2 c3'). *) - - fun unfold_defs thy t = +fun unfold_defs thy t = let (* Term.term -> Term.term *) fun unfold_loop t = @@ -658,13 +615,13 @@ | Const (@{const_name Finite_Set.card}, _) => t | Const (@{const_name Finite_Set.finite}, _) => t | Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat}, - Type ("fun", [@{typ nat}, @{typ bool}])])) => t + Type ("fun", [@{typ nat}, @{typ bool}])])) => t | Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat}, - Type ("fun", [@{typ nat}, @{typ nat}])])) => t + Type ("fun", [@{typ nat}, @{typ nat}])])) => t | Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat}, - Type ("fun", [@{typ nat}, @{typ nat}])])) => t + Type ("fun", [@{typ nat}, @{typ nat}])])) => t | Const (@{const_name Groups.times}, Type ("fun", [@{typ nat}, - Type ("fun", [@{typ nat}, @{typ nat}])])) => t + Type ("fun", [@{typ nat}, @{typ nat}])])) => t | Const (@{const_name List.append}, _) => t (* UNSOUND | Const (@{const_name lfp}, _) => t @@ -674,27 +631,28 @@ | Const (@{const_name snd}, _) => t (* simply-typed lambda calculus *) | Const (s, T) => - (if is_IDT_constructor thy (s, T) - orelse is_IDT_recursor thy (s, T) then - t (* do not unfold IDT constructors/recursors *) - (* unfold the constant if there is a defining equation *) - else case get_def thy (s, T) of - SOME (axname, rhs) => - (* Note: if the term to be unfolded (i.e. 'Const (s, T)') *) - (* occurs on the right-hand side of the equation, i.e. in *) - (* 'rhs', we must not use this equation to unfold, because *) - (* that would loop. Here would be the right place to *) - (* check this. However, getting this really right seems *) - (* difficult because the user may state arbitrary axioms, *) - (* which could interact with overloading to create loops. *) - ((*tracing (" unfolding: " ^ axname);*) - unfold_loop rhs) - | NONE => t) - | Free _ => t - | Var _ => t - | Bound _ => t + (if is_IDT_constructor thy (s, T) + orelse is_IDT_recursor thy (s, T) then + t (* do not unfold IDT constructors/recursors *) + (* unfold the constant if there is a defining equation *) + else + case get_def thy (s, T) of + SOME (axname, rhs) => + (* Note: if the term to be unfolded (i.e. 'Const (s, T)') *) + (* occurs on the right-hand side of the equation, i.e. in *) + (* 'rhs', we must not use this equation to unfold, because *) + (* that would loop. Here would be the right place to *) + (* check this. However, getting this really right seems *) + (* difficult because the user may state arbitrary axioms, *) + (* which could interact with overloading to create loops. *) + ((*tracing (" unfolding: " ^ axname);*) + unfold_loop rhs) + | NONE => t) + | Free _ => t + | Var _ => t + | Bound _ => t | Abs (s, T, body) => Abs (s, T, unfold_loop body) - | t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2) + | t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2) val result = Envir.beta_eta_contract (unfold_loop t) in result @@ -705,28 +663,27 @@ (* versions of) all HOL axioms that are relevant w.r.t 't' *) (* ------------------------------------------------------------------------- *) - (* Note: to make the collection of axioms more easily extensible, this *) - (* function could be based on user-supplied "axiom collectors", *) - (* similar to 'interpret'/interpreters or 'print'/printers *) +(* Note: to make the collection of axioms more easily extensible, this *) +(* function could be based on user-supplied "axiom collectors", *) +(* similar to 'interpret'/interpreters or 'print'/printers *) - (* Note: currently we use "inverse" functions to the definitional *) - (* mechanisms provided by Isabelle/HOL, e.g. for "axclass", *) - (* "typedef", "definition". A more general approach could consider *) - (* *every* axiom of the theory and collect it if it has a constant/ *) - (* type/typeclass in common with the term 't'. *) +(* Note: currently we use "inverse" functions to the definitional *) +(* mechanisms provided by Isabelle/HOL, e.g. for "axclass", *) +(* "typedef", "definition". A more general approach could consider *) +(* *every* axiom of the theory and collect it if it has a constant/ *) +(* type/typeclass in common with the term 't'. *) - (* theory -> Term.term -> Term.term list *) +(* Which axioms are "relevant" for a particular term/type goes hand in *) +(* hand with the interpretation of that term/type by its interpreter (see *) +(* way below): if the interpretation respects an axiom anyway, the axiom *) +(* does not need to be added as a constraint here. *) - (* Which axioms are "relevant" for a particular term/type goes hand in *) - (* hand with the interpretation of that term/type by its interpreter (see *) - (* way below): if the interpretation respects an axiom anyway, the axiom *) - (* does not need to be added as a constraint here. *) +(* To avoid collecting the same axiom multiple times, we use an *) +(* accumulator 'axs' which contains all axioms collected so far. *) - (* To avoid collecting the same axiom multiple times, we use an *) - (* accumulator 'axs' which contains all axioms collected so far. *) - - fun collect_axioms thy t = +fun collect_axioms ctxt t = let + val thy = ProofContext.theory_of ctxt val _ = tracing "Adding axioms..." val axioms = Theory.all_axioms_of thy fun collect_this_axiom (axname, ax) axs = @@ -743,7 +700,7 @@ TFree (_, sort) => sort | TVar (_, sort) => sort | _ => raise REFUTE ("collect_axioms", - "type " ^ Syntax.string_of_typ_global thy T ^ " is not a variable")) + "type " ^ Syntax.string_of_typ ctxt T ^ " is not a variable")) (* obtain axioms for all superclasses *) val superclasses = sort @ maps (Sign.super_classes thy) sort (* merely an optimization, because 'collect_this_axiom' disallows *) @@ -761,7 +718,7 @@ | [(idx, S)] => (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax) | _ => raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^ - Syntax.string_of_term_global thy ax ^ + Syntax.string_of_term ctxt ax ^ ") contains more than one type variable"))) class_axioms in @@ -782,7 +739,7 @@ | NONE => (case get_typedef thy T of SOME (axname, ax) => - collect_this_axiom (axname, ax) axs + collect_this_axiom (axname, ax) axs | NONE => (* unspecified type, perhaps introduced with "typedecl" *) (* at least collect relevant type axioms for the argument types *) @@ -808,19 +765,19 @@ | Const (@{const_name False}, _) => axs | Const (@{const_name undefined}, T) => collect_type_axioms T axs | Const (@{const_name The}, T) => - let - val ax = specialize_type thy (@{const_name The}, T) - (the (AList.lookup (op =) axioms "HOL.the_eq_trivial")) - in - collect_this_axiom ("HOL.the_eq_trivial", ax) axs - end + let + val ax = specialize_type thy (@{const_name The}, T) + (the (AList.lookup (op =) axioms "HOL.the_eq_trivial")) + in + collect_this_axiom ("HOL.the_eq_trivial", ax) axs + end | Const (@{const_name Hilbert_Choice.Eps}, T) => - let - val ax = specialize_type thy (@{const_name Hilbert_Choice.Eps}, T) - (the (AList.lookup (op =) axioms "Hilbert_Choice.someI")) - in - collect_this_axiom ("Hilbert_Choice.someI", ax) axs - end + let + val ax = specialize_type thy (@{const_name Hilbert_Choice.Eps}, T) + (the (AList.lookup (op =) axioms "Hilbert_Choice.someI")) + in + collect_this_axiom ("Hilbert_Choice.someI", ax) axs + end | Const (@{const_name All}, T) => collect_type_axioms T axs | Const (@{const_name Ex}, T) => collect_type_axioms T axs | Const (@{const_name HOL.eq}, T) => collect_type_axioms T axs @@ -864,13 +821,14 @@ val ax_in = SOME (specialize_type thy (s, T) of_class) (* type match may fail due to sort constraints *) handle Type.TYPE_MATCH => NONE - val ax_1 = Option.map (fn ax => (Syntax.string_of_term_global thy ax, ax)) ax_in + val ax_1 = Option.map (fn ax => (Syntax.string_of_term ctxt ax, ax)) ax_in val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class) in collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs) end else if is_IDT_constructor thy (s, T) - orelse is_IDT_recursor thy (s, T) then + orelse is_IDT_recursor thy (s, T) + then (* only collect relevant type axioms *) collect_type_axioms T axs else @@ -898,70 +856,71 @@ (* and all mutually recursive IDTs are considered. *) (* ------------------------------------------------------------------------- *) - fun ground_types thy t = +fun ground_types ctxt t = let + val thy = ProofContext.theory_of ctxt fun collect_types T acc = (case T of Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc) - | Type ("prop", []) => acc - | Type (s, Ts) => - (case Datatype.get_info thy s of - SOME info => (* inductive datatype *) - let - val index = #index info - val descr = #descr info - val (_, typs, _) = the (AList.lookup (op =) descr index) - val typ_assoc = typs ~~ Ts - (* sanity check: every element in 'dtyps' must be a *) - (* 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then - raise REFUTE ("ground_types", "datatype argument (for type " - ^ Syntax.string_of_typ_global thy T ^ ") is not a variable") - else () - (* required for mutually recursive datatypes; those need to *) - (* be added even if they are an instance of an otherwise non- *) - (* recursive datatype *) - fun collect_dtyp d acc = - let - val dT = typ_of_dtyp descr typ_assoc d - in - case d of - Datatype_Aux.DtTFree _ => - collect_types dT acc - | Datatype_Aux.DtType (_, ds) => - collect_types dT (fold_rev collect_dtyp ds acc) - | Datatype_Aux.DtRec i => - if member (op =) acc dT then - acc (* prevent infinite recursion *) - else + | Type ("prop", []) => acc + | Type (s, Ts) => + (case Datatype.get_info thy s of + SOME info => (* inductive datatype *) + let + val index = #index info + val descr = #descr info + val (_, typs, _) = the (AList.lookup (op =) descr index) + val typ_assoc = typs ~~ Ts + (* sanity check: every element in 'dtyps' must be a *) + (* 'DtTFree' *) + val _ = if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then + raise REFUTE ("ground_types", "datatype argument (for type " + ^ Syntax.string_of_typ ctxt T ^ ") is not a variable") + else () + (* required for mutually recursive datatypes; those need to *) + (* be added even if they are an instance of an otherwise non- *) + (* recursive datatype *) + fun collect_dtyp d acc = let - val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i) - (* if the current type is a recursive IDT (i.e. a depth *) - (* is required), add it to 'acc' *) - val acc_dT = if Library.exists (fn (_, ds) => - Library.exists Datatype_Aux.is_rec_type ds) dconstrs then - insert (op =) dT acc - else acc - (* collect argument types *) - val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT - (* collect constructor types *) - val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps + val dT = typ_of_dtyp descr typ_assoc d in - acc_dconstrs + case d of + Datatype_Aux.DtTFree _ => + collect_types dT acc + | Datatype_Aux.DtType (_, ds) => + collect_types dT (fold_rev collect_dtyp ds acc) + | Datatype_Aux.DtRec i => + if member (op =) acc dT then + acc (* prevent infinite recursion *) + else + let + val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i) + (* if the current type is a recursive IDT (i.e. a depth *) + (* is required), add it to 'acc' *) + val acc_dT = if Library.exists (fn (_, ds) => + Library.exists Datatype_Aux.is_rec_type ds) dconstrs then + insert (op =) dT acc + else acc + (* collect argument types *) + val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT + (* collect constructor types *) + val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps + in + acc_dconstrs + end end - end - in - (* argument types 'Ts' could be added here, but they are also *) - (* added by 'collect_dtyp' automatically *) - collect_dtyp (Datatype_Aux.DtRec index) acc - end - | NONE => - (* not an inductive datatype, e.g. defined via "typedef" or *) - (* "typedecl" *) - insert (op =) T (fold collect_types Ts acc)) - | TFree _ => insert (op =) T acc - | TVar _ => insert (op =) T acc) + in + (* argument types 'Ts' could be added here, but they are also *) + (* added by 'collect_dtyp' automatically *) + collect_dtyp (Datatype_Aux.DtRec index) acc + end + | NONE => + (* not an inductive datatype, e.g. defined via "typedef" or *) + (* "typedecl" *) + insert (op =) T (fold collect_types Ts acc)) + | TFree _ => insert (op =) T acc + | TVar _ => insert (op =) T acc) in fold_types collect_types t [] end; @@ -973,11 +932,11 @@ (* list") are identified. *) (* ------------------------------------------------------------------------- *) - (* Term.typ -> string *) +(* Term.typ -> string *) - fun string_of_typ (Type (s, _)) = s - | string_of_typ (TFree (s, _)) = s - | string_of_typ (TVar ((s,_), _)) = s; +fun string_of_typ (Type (s, _)) = s + | string_of_typ (TFree (s, _)) = s + | string_of_typ (TVar ((s,_), _)) = s; (* ------------------------------------------------------------------------- *) (* first_universe: returns the "first" (i.e. smallest) universe by assigning *) @@ -985,9 +944,9 @@ (* 'sizes' *) (* ------------------------------------------------------------------------- *) - (* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *) +(* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *) - fun first_universe xs sizes minsize = +fun first_universe xs sizes minsize = let fun size_of_typ T = case AList.lookup (op =) sizes (string_of_typ T) of @@ -1004,39 +963,39 @@ (* type may have a fixed size given in 'sizes' *) (* ------------------------------------------------------------------------- *) - (* (Term.typ * int) list -> (string * int) list -> int -> int -> - (Term.typ * int) list option *) +(* (Term.typ * int) list -> (string * int) list -> int -> int -> + (Term.typ * int) list option *) - fun next_universe xs sizes minsize maxsize = +fun next_universe xs sizes minsize maxsize = let (* creates the "first" list of length 'len', where the sum of all list *) (* elements is 'sum', and the length of the list is 'len' *) (* int -> int -> int -> int list option *) fun make_first _ 0 sum = - if sum=0 then - SOME [] - else - NONE + if sum = 0 then + SOME [] + else + NONE | make_first max len sum = - if sum<=max orelse max<0 then - Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0) - else - Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max)) + if sum <= max orelse max < 0 then + Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0) + else + Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max)) (* enumerates all int lists with a fixed length, where 0<=x<='max' for *) (* all list elements x (unless 'max'<0) *) (* int -> int -> int -> int list -> int list option *) fun next max len sum [] = - NONE + NONE | next max len sum [x] = - (* we've reached the last list element, so there's no shift possible *) - make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *) + (* we've reached the last list element, so there's no shift possible *) + make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *) | next max len sum (x1::x2::xs) = - if x1>0 andalso (x2<max orelse max<0) then - (* we can shift *) - SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs) - else - (* continue search *) - next max (len+1) (sum+x1) (x2::xs) + if x1>0 andalso (x2<max orelse max<0) then + (* we can shift *) + SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs) + else + (* continue search *) + next max (len+1) (sum+x1) (x2::xs) (* only consider those types for which the size is not fixed *) val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs (* subtract 'minsize' from every size (will be added again at the end) *) @@ -1044,16 +1003,15 @@ in case next (maxsize-minsize) 0 0 diffs of SOME diffs' => - (* merge with those types for which the size is fixed *) - SOME (fst (fold_map (fn (T, _) => fn ds => - case AList.lookup (op =) sizes (string_of_typ T) of - (* return the fixed size *) - SOME n => ((T, n), ds) - (* consume the head of 'ds', add 'minsize' *) - | NONE => ((T, minsize + hd ds), tl ds)) - xs diffs')) - | NONE => - NONE + (* merge with those types for which the size is fixed *) + SOME (fst (fold_map (fn (T, _) => fn ds => + case AList.lookup (op =) sizes (string_of_typ T) of + (* return the fixed size *) + SOME n => ((T, n), ds) + (* consume the head of 'ds', add 'minsize' *) + | NONE => ((T, minsize + hd ds), tl ds)) + xs diffs')) + | NONE => NONE end; (* ------------------------------------------------------------------------- *) @@ -1061,12 +1019,10 @@ (* formula that is true iff the interpretation denotes "true" *) (* ------------------------------------------------------------------------- *) - (* interpretation -> prop_formula *) +(* interpretation -> prop_formula *) - fun toTrue (Leaf [fm, _]) = - fm - | toTrue _ = - raise REFUTE ("toTrue", "interpretation does not denote a Boolean value"); +fun toTrue (Leaf [fm, _]) = fm + | toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value"); (* ------------------------------------------------------------------------- *) (* toFalse: converts the interpretation of a Boolean value to a *) @@ -1074,170 +1030,168 @@ (* denotes "false" *) (* ------------------------------------------------------------------------- *) - (* interpretation -> prop_formula *) +(* interpretation -> prop_formula *) - fun toFalse (Leaf [_, fm]) = - fm - | toFalse _ = - raise REFUTE ("toFalse", "interpretation does not denote a Boolean value"); +fun toFalse (Leaf [_, fm]) = fm + | toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value"); (* ------------------------------------------------------------------------- *) (* find_model: repeatedly calls 'interpret' with appropriate parameters, *) (* applies a SAT solver, and (in case a model is found) displays *) (* the model to the user by calling 'print_model' *) -(* thy : the current theory *) (* {...} : parameters that control the translation/model generation *) (* assm_ts : assumptions to be considered unless "no_assms" is specified *) (* t : term to be translated into a propositional formula *) (* negate : if true, find a model that makes 't' false (rather than true) *) (* ------------------------------------------------------------------------- *) - (* theory -> params -> Term.term -> bool -> unit *) - - fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver, - no_assms, expect} assm_ts t negate = +fun find_model ctxt + {sizes, minsize, maxsize, maxvars, maxtime, satsolver, no_assms, expect} + assm_ts t negate = let + val thy = ProofContext.theory_of ctxt (* string -> unit *) fun check_expect outcome_code = if expect = "" orelse outcome_code = expect then () else error ("Unexpected outcome: " ^ quote outcome_code ^ ".") (* unit -> unit *) fun wrapper () = - let - val timer = Timer.startRealTimer () - val t = if no_assms then t - else if negate then Logic.list_implies (assm_ts, t) - else Logic.mk_conjunction_list (t :: assm_ts) - val u = unfold_defs thy t - val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term_global thy u) - val axioms = collect_axioms thy u - (* Term.typ list *) - val types = fold (union (op =) o ground_types thy) (u :: axioms) [] - val _ = tracing ("Ground types: " - ^ (if null types then "none." - else commas (map (Syntax.string_of_typ_global thy) types))) - (* we can only consider fragments of recursive IDTs, so we issue a *) - (* warning if the formula contains a recursive IDT *) - (* TODO: no warning needed for /positive/ occurrences of IDTs *) - val maybe_spurious = Library.exists (fn - Type (s, _) => - (case Datatype.get_info thy s of - SOME info => (* inductive datatype *) - let - val index = #index info - val descr = #descr info - val (_, _, constrs) = the (AList.lookup (op =) descr index) - in - (* recursive datatype? *) - Library.exists (fn (_, ds) => - Library.exists Datatype_Aux.is_rec_type ds) constrs - end - | NONE => false) - | _ => false) types - val _ = if maybe_spurious then - warning ("Term contains a recursive datatype; " - ^ "countermodel(s) may be spurious!") - else - () - (* (Term.typ * int) list -> string *) - fun find_model_loop universe = let - val msecs_spent = Time.toMilliseconds (Timer.checkRealTimer timer) - val _ = maxtime = 0 orelse msecs_spent < 1000 * maxtime - orelse raise TimeLimit.TimeOut - val init_model = (universe, []) - val init_args = {maxvars = maxvars, def_eq = false, next_idx = 1, - bounds = [], wellformed = True} - val _ = tracing ("Translating term (sizes: " - ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...") - (* translate 'u' and all axioms *) - val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) => + val timer = Timer.startRealTimer () + val t = + if no_assms then t + else if negate then Logic.list_implies (assm_ts, t) + else Logic.mk_conjunction_list (t :: assm_ts) + val u = unfold_defs thy t + val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term ctxt u) + val axioms = collect_axioms ctxt u + (* Term.typ list *) + val types = fold (union (op =) o ground_types ctxt) (u :: axioms) [] + val _ = tracing ("Ground types: " + ^ (if null types then "none." + else commas (map (Syntax.string_of_typ ctxt) types))) + (* we can only consider fragments of recursive IDTs, so we issue a *) + (* warning if the formula contains a recursive IDT *) + (* TODO: no warning needed for /positive/ occurrences of IDTs *) + val maybe_spurious = Library.exists (fn + Type (s, _) => + (case Datatype.get_info thy s of + SOME info => (* inductive datatype *) + let + val index = #index info + val descr = #descr info + val (_, _, constrs) = the (AList.lookup (op =) descr index) + in + (* recursive datatype? *) + Library.exists (fn (_, ds) => + Library.exists Datatype_Aux.is_rec_type ds) constrs + end + | NONE => false) + | _ => false) types + val _ = + if maybe_spurious then + warning ("Term contains a recursive datatype; " + ^ "countermodel(s) may be spurious!") + else + () + (* (Term.typ * int) list -> string *) + fun find_model_loop universe = let - val (i, m', a') = interpret thy m a t' + val msecs_spent = Time.toMilliseconds (Timer.checkRealTimer timer) + val _ = maxtime = 0 orelse msecs_spent < 1000 * maxtime + orelse raise TimeLimit.TimeOut + val init_model = (universe, []) + val init_args = {maxvars = maxvars, def_eq = false, next_idx = 1, + bounds = [], wellformed = True} + val _ = tracing ("Translating term (sizes: " + ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...") + (* translate 'u' and all axioms *) + val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) => + let + val (i, m', a') = interpret ctxt m a t' + in + (* set 'def_eq' to 'true' *) + (i, (m', {maxvars = #maxvars a', def_eq = true, + next_idx = #next_idx a', bounds = #bounds a', + wellformed = #wellformed a'})) + end) (u :: axioms) (init_model, init_args) + (* make 'u' either true or false, and make all axioms true, and *) + (* add the well-formedness side condition *) + val fm_u = (if negate then toFalse else toTrue) (hd intrs) + val fm_ax = PropLogic.all (map toTrue (tl intrs)) + val fm = PropLogic.all [#wellformed args, fm_ax, fm_u] + val _ = + (if satsolver = "dpll" orelse satsolver = "enumerate" then + warning ("Using SAT solver " ^ quote satsolver ^ + "; for better performance, consider installing an \ + \external solver.") + else ()); + val solver = + SatSolver.invoke_solver satsolver + handle Option.Option => + error ("Unknown SAT solver: " ^ quote satsolver ^ + ". Available solvers: " ^ + commas (map (quote o fst) (!SatSolver.solvers)) ^ ".") in - (* set 'def_eq' to 'true' *) - (i, (m', {maxvars = #maxvars a', def_eq = true, - next_idx = #next_idx a', bounds = #bounds a', - wellformed = #wellformed a'})) - end) (u :: axioms) (init_model, init_args) - (* make 'u' either true or false, and make all axioms true, and *) - (* add the well-formedness side condition *) - val fm_u = (if negate then toFalse else toTrue) (hd intrs) - val fm_ax = PropLogic.all (map toTrue (tl intrs)) - val fm = PropLogic.all [#wellformed args, fm_ax, fm_u] - val _ = - (if satsolver = "dpll" orelse satsolver = "enumerate" then - warning ("Using SAT solver " ^ quote satsolver ^ - "; for better performance, consider installing an \ - \external solver.") - else - ()); - val solver = - SatSolver.invoke_solver satsolver - handle Option.Option => - error ("Unknown SAT solver: " ^ quote satsolver ^ - ". Available solvers: " ^ - commas (map (quote o fst) (!SatSolver.solvers)) ^ ".") - in - priority "Invoking SAT solver..."; - (case solver fm of - SatSolver.SATISFIABLE assignment => - (priority ("*** Model found: ***\n" ^ print_model thy model - (fn i => case assignment i of SOME b => b | NONE => true)); - if maybe_spurious then "potential" else "genuine") - | SatSolver.UNSATISFIABLE _ => - (priority "No model exists."; - case next_universe universe sizes minsize maxsize of - SOME universe' => find_model_loop universe' - | NONE => (priority - "Search terminated, no larger universe within the given limits."; - "none")) - | SatSolver.UNKNOWN => - (priority "No model found."; - case next_universe universe sizes minsize maxsize of - SOME universe' => find_model_loop universe' - | NONE => (priority - "Search terminated, no larger universe within the given limits."; - "unknown")) - ) handle SatSolver.NOT_CONFIGURED => - (error ("SAT solver " ^ quote satsolver ^ " is not configured."); - "unknown") - end handle MAXVARS_EXCEEDED => - (priority ("Search terminated, number of Boolean variables (" - ^ string_of_int maxvars ^ " allowed) exceeded."); - "unknown") + priority "Invoking SAT solver..."; + (case solver fm of + SatSolver.SATISFIABLE assignment => + (priority ("*** Model found: ***\n" ^ print_model ctxt model + (fn i => case assignment i of SOME b => b | NONE => true)); + if maybe_spurious then "potential" else "genuine") + | SatSolver.UNSATISFIABLE _ => + (priority "No model exists."; + case next_universe universe sizes minsize maxsize of + SOME universe' => find_model_loop universe' + | NONE => (priority + "Search terminated, no larger universe within the given limits."; + "none")) + | SatSolver.UNKNOWN => + (priority "No model found."; + case next_universe universe sizes minsize maxsize of + SOME universe' => find_model_loop universe' + | NONE => (priority + "Search terminated, no larger universe within the given limits."; + "unknown"))) handle SatSolver.NOT_CONFIGURED => + (error ("SAT solver " ^ quote satsolver ^ " is not configured."); + "unknown") + end + handle MAXVARS_EXCEEDED => + (priority ("Search terminated, number of Boolean variables (" + ^ string_of_int maxvars ^ " allowed) exceeded."); + "unknown") + val outcome_code = find_model_loop (first_universe types sizes minsize) in check_expect outcome_code end - in - (* some parameter sanity checks *) - minsize>=1 orelse - error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1"); - maxsize>=1 orelse - error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1"); - maxsize>=minsize orelse - error ("\"maxsize\" (=" ^ string_of_int maxsize ^ - ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ")."); - maxvars>=0 orelse - error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0"); - maxtime>=0 orelse - error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0"); - (* enter loop with or without time limit *) - priority ("Trying to find a model that " - ^ (if negate then "refutes" else "satisfies") ^ ": " - ^ Syntax.string_of_term_global thy t); - if maxtime>0 then ( - TimeLimit.timeLimit (Time.fromSeconds maxtime) - wrapper () - handle TimeLimit.TimeOut => - (priority ("Search terminated, time limit (" ^ - string_of_int maxtime - ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded."); - check_expect "unknown") - ) else + in + (* some parameter sanity checks *) + minsize>=1 orelse + error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1"); + maxsize>=1 orelse + error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1"); + maxsize>=minsize orelse + error ("\"maxsize\" (=" ^ string_of_int maxsize ^ + ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ")."); + maxvars>=0 orelse + error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0"); + maxtime>=0 orelse + error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0"); + (* enter loop with or without time limit *) + priority ("Trying to find a model that " + ^ (if negate then "refutes" else "satisfies") ^ ": " + ^ Syntax.string_of_term ctxt t); + if maxtime > 0 then ( + TimeLimit.timeLimit (Time.fromSeconds maxtime) wrapper () - end; + handle TimeLimit.TimeOut => + (priority ("Search terminated, time limit (" ^ + string_of_int maxtime + ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded."); + check_expect "unknown") + ) else wrapper () + end; (* ------------------------------------------------------------------------- *) @@ -1250,10 +1204,8 @@ (* parameters *) (* ------------------------------------------------------------------------- *) - (* theory -> (string * string) list -> Term.term list -> Term.term -> unit *) - - fun satisfy_term thy params assm_ts t = - find_model thy (actual_params thy params) assm_ts t false; +fun satisfy_term ctxt params assm_ts t = + find_model ctxt (actual_params ctxt params) assm_ts t false; (* ------------------------------------------------------------------------- *) (* refute_term: calls 'find_model' to find a model that refutes 't' *) @@ -1261,9 +1213,7 @@ (* parameters *) (* ------------------------------------------------------------------------- *) - (* theory -> (string * string) list -> Term.term list -> Term.term -> unit *) - - fun refute_term thy params assm_ts t = +fun refute_term ctxt params assm_ts t = let (* disallow schematic type variables, since we cannot properly negate *) (* terms containing them (their logical meaning is that there EXISTS a *) @@ -1293,33 +1243,32 @@ (* quantified variables. *) (* maps !!x1...xn. !xk...xm. t to t *) fun strip_all_body (Const (@{const_name all}, _) $ Abs (_, _, t)) = - strip_all_body t + strip_all_body t | strip_all_body (Const (@{const_name Trueprop}, _) $ t) = - strip_all_body t + strip_all_body t | strip_all_body (Const (@{const_name All}, _) $ Abs (_, _, t)) = - strip_all_body t + strip_all_body t | strip_all_body t = t (* maps !!x1...xn. !xk...xm. t to [x1, ..., xn, xk, ..., xm] *) fun strip_all_vars (Const (@{const_name all}, _) $ Abs (a, T, t)) = - (a, T) :: strip_all_vars t + (a, T) :: strip_all_vars t | strip_all_vars (Const (@{const_name Trueprop}, _) $ t) = - strip_all_vars t + strip_all_vars t | strip_all_vars (Const (@{const_name All}, _) $ Abs (a, T, t)) = - (a, T) :: strip_all_vars t - | strip_all_vars t = - [] : (string * typ) list + (a, T) :: strip_all_vars t + | strip_all_vars t = [] : (string * typ) list val strip_t = strip_all_body ex_closure - val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure) + val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure) val subst_t = Term.subst_bounds (map Free frees, strip_t) in - find_model thy (actual_params thy params) assm_ts subst_t true + find_model ctxt (actual_params ctxt params) assm_ts subst_t true end; (* ------------------------------------------------------------------------- *) (* refute_goal *) (* ------------------------------------------------------------------------- *) - fun refute_goal ctxt params th i = +fun refute_goal ctxt params th i = let val t = th |> prop_of in @@ -1330,8 +1279,7 @@ val assms = map term_of (Assumption.all_assms_of ctxt) val (t, frees) = Logic.goal_params t i in - refute_term (ProofContext.theory_of ctxt) params assms - (subst_bounds (frees, t)) + refute_term ctxt params assms (subst_bounds (frees, t)) end end @@ -1346,81 +1294,64 @@ (* variables) *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> Term.typ -> interpretation list *) - - fun make_constants thy model T = +fun make_constants ctxt model T = let (* returns a list with all unit vectors of length n *) (* int -> interpretation list *) fun unit_vectors n = - let - (* returns the k-th unit vector of length n *) - (* int * int -> interpretation *) - fun unit_vector (k, n) = - Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False))) - (* int -> interpretation list *) - fun unit_vectors_loop k = - if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1) - in - unit_vectors_loop 1 - end + let + (* returns the k-th unit vector of length n *) + (* int * int -> interpretation *) + fun unit_vector (k, n) = + Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False))) + (* int -> interpretation list *) + fun unit_vectors_loop k = + if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1) + in + unit_vectors_loop 1 + end (* returns a list of lists, each one consisting of n (possibly *) (* identical) elements from 'xs' *) (* int -> 'a list -> 'a list list *) - fun pick_all 1 xs = - map single xs + fun pick_all 1 xs = map single xs | pick_all n xs = - let val rec_pick = pick_all (n-1) xs in - maps (fn x => map (cons x) rec_pick) xs - end + let val rec_pick = pick_all (n - 1) xs in + maps (fn x => map (cons x) rec_pick) xs + end (* returns all constant interpretations that have the same tree *) (* structure as the interpretation argument *) (* interpretation -> interpretation list *) fun make_constants_intr (Leaf xs) = unit_vectors (length xs) | make_constants_intr (Node xs) = map Node (pick_all (length xs) - (make_constants_intr (hd xs))) + (make_constants_intr (hd xs))) (* obtain the interpretation for a variable of type 'T' *) - val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, + val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T)) in make_constants_intr i end; (* ------------------------------------------------------------------------- *) -(* power: 'power (a, b)' computes a^b, for a>=0, b>=0 *) -(* ------------------------------------------------------------------------- *) - - (* int * int -> int *) - - fun power (a, 0) = 1 - | power (a, 1) = a - | power (a, b) = let val ab = power(a, b div 2) in - ab * ab * power(a, b mod 2) - end; - -(* ------------------------------------------------------------------------- *) (* size_of_type: returns the number of elements in a type 'T' (i.e. 'length *) (* (make_constants T)', but implemented more efficiently) *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> Term.typ -> int *) +(* returns 0 for an empty ground type or a function type with empty *) +(* codomain, but fails for a function type with empty domain -- *) +(* admissibility of datatype constructor argument types (see "Inductive *) +(* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel, *) +(* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *) +(* never occur as the domain of a function type that is the type of a *) +(* constructor argument *) - (* returns 0 for an empty ground type or a function type with empty *) - (* codomain, but fails for a function type with empty domain -- *) - (* admissibility of datatype constructor argument types (see "Inductive *) - (* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel, *) - (* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *) - (* never occur as the domain of a function type that is the type of a *) - (* constructor argument *) - - fun size_of_type thy model T = +fun size_of_type ctxt model T = let (* returns the number of elements that have the same tree structure as a *) (* given interpretation *) fun size_of_intr (Leaf xs) = length xs - | size_of_intr (Node xs) = power (size_of_intr (hd xs), length xs) + | size_of_intr (Node xs) = Integer.pow (length xs) (size_of_intr (hd xs)) (* obtain the interpretation for a variable of type 'T' *) - val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, + val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T)) in size_of_intr i @@ -1430,11 +1361,11 @@ (* TT/FF: interpretations that denote "true" or "false", respectively *) (* ------------------------------------------------------------------------- *) - (* interpretation *) +(* interpretation *) - val TT = Leaf [True, False]; +val TT = Leaf [True, False]; - val FF = Leaf [False, True]; +val FF = Leaf [False, True]; (* ------------------------------------------------------------------------- *) (* make_equality: returns an interpretation that denotes (extensional) *) @@ -1447,46 +1378,46 @@ (* 'not_equal' to another interpretation *) (* ------------------------------------------------------------------------- *) - (* We could in principle represent '=' on a type T by a particular *) - (* interpretation. However, the size of that interpretation is quadratic *) - (* in the size of T. Therefore comparing the interpretations 'i1' and *) - (* 'i2' directly is more efficient than constructing the interpretation *) - (* for equality on T first, and "applying" this interpretation to 'i1' *) - (* and 'i2' in the usual way (cf. 'interpretation_apply') then. *) +(* We could in principle represent '=' on a type T by a particular *) +(* interpretation. However, the size of that interpretation is quadratic *) +(* in the size of T. Therefore comparing the interpretations 'i1' and *) +(* 'i2' directly is more efficient than constructing the interpretation *) +(* for equality on T first, and "applying" this interpretation to 'i1' *) +(* and 'i2' in the usual way (cf. 'interpretation_apply') then. *) - (* interpretation * interpretation -> interpretation *) +(* interpretation * interpretation -> interpretation *) - fun make_equality (i1, i2) = +fun make_equality (i1, i2) = let (* interpretation * interpretation -> prop_formula *) fun equal (i1, i2) = (case i1 of Leaf xs => - (case i2 of - Leaf ys => PropLogic.dot_product (xs, ys) (* defined and equal *) - | Node _ => raise REFUTE ("make_equality", - "second interpretation is higher")) + (case i2 of + Leaf ys => PropLogic.dot_product (xs, ys) (* defined and equal *) + | Node _ => raise REFUTE ("make_equality", + "second interpretation is higher")) | Node xs => - (case i2 of - Leaf _ => raise REFUTE ("make_equality", - "first interpretation is higher") - | Node ys => PropLogic.all (map equal (xs ~~ ys)))) + (case i2 of + Leaf _ => raise REFUTE ("make_equality", + "first interpretation is higher") + | Node ys => PropLogic.all (map equal (xs ~~ ys)))) (* interpretation * interpretation -> prop_formula *) fun not_equal (i1, i2) = (case i1 of Leaf xs => - (case i2 of - (* defined and not equal *) - Leaf ys => PropLogic.all ((PropLogic.exists xs) - :: (PropLogic.exists ys) - :: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys))) - | Node _ => raise REFUTE ("make_equality", - "second interpretation is higher")) + (case i2 of + (* defined and not equal *) + Leaf ys => PropLogic.all ((PropLogic.exists xs) + :: (PropLogic.exists ys) + :: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys))) + | Node _ => raise REFUTE ("make_equality", + "second interpretation is higher")) | Node xs => - (case i2 of - Leaf _ => raise REFUTE ("make_equality", - "first interpretation is higher") - | Node ys => PropLogic.exists (map not_equal (xs ~~ ys)))) + (case i2 of + Leaf _ => raise REFUTE ("make_equality", + "first interpretation is higher") + | Node ys => PropLogic.exists (map not_equal (xs ~~ ys)))) in (* a value may be undefined; therefore 'not_equal' is not just the *) (* negation of 'equal' *) @@ -1502,25 +1433,25 @@ (* to an undefined interpretation. *) (* ------------------------------------------------------------------------- *) - (* interpretation * interpretation -> interpretation *) +(* interpretation * interpretation -> interpretation *) - fun make_def_equality (i1, i2) = +fun make_def_equality (i1, i2) = let (* interpretation * interpretation -> prop_formula *) fun equal (i1, i2) = (case i1 of Leaf xs => - (case i2 of - (* defined and equal, or both undefined *) - Leaf ys => SOr (PropLogic.dot_product (xs, ys), - SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys))) - | Node _ => raise REFUTE ("make_def_equality", - "second interpretation is higher")) + (case i2 of + (* defined and equal, or both undefined *) + Leaf ys => SOr (PropLogic.dot_product (xs, ys), + SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys))) + | Node _ => raise REFUTE ("make_def_equality", + "second interpretation is higher")) | Node xs => - (case i2 of - Leaf _ => raise REFUTE ("make_def_equality", - "first interpretation is higher") - | Node ys => PropLogic.all (map equal (xs ~~ ys)))) + (case i2 of + Leaf _ => raise REFUTE ("make_def_equality", + "first interpretation is higher") + | Node ys => PropLogic.all (map equal (xs ~~ ys)))) (* interpretation *) val eq = equal (i1, i2) in @@ -1533,9 +1464,9 @@ (* argument denoted by 'i2' *) (* ------------------------------------------------------------------------- *) - (* interpretation * interpretation -> interpretation *) +(* interpretation * interpretation -> interpretation *) - fun interpretation_apply (i1, i2) = +fun interpretation_apply (i1, i2) = let (* interpretation * interpretation -> interpretation *) fun interpretation_disjunction (tr1,tr2) = @@ -1546,50 +1477,46 @@ tree_map (map (fn x => SAnd (fm,x))) tr (* prop_formula list * interpretation list -> interpretation *) fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) = - prop_formula_times_interpretation (fm,tr) + prop_formula_times_interpretation (fm,tr) | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) = - interpretation_disjunction (prop_formula_times_interpretation (fm,tr), - prop_formula_list_dot_product_interpretation_list (fms,trees)) + interpretation_disjunction (prop_formula_times_interpretation (fm,tr), + prop_formula_list_dot_product_interpretation_list (fms,trees)) | prop_formula_list_dot_product_interpretation_list (_,_) = - raise REFUTE ("interpretation_apply", "empty list (in dot product)") + raise REFUTE ("interpretation_apply", "empty list (in dot product)") (* concatenates 'x' with every list in 'xss', returning a new list of *) (* lists *) (* 'a -> 'a list list -> 'a list list *) - fun cons_list x xss = - map (cons x) xss + fun cons_list x xss = map (cons x) xss (* returns a list of lists, each one consisting of one element from each *) (* element of 'xss' *) (* 'a list list -> 'a list list *) - fun pick_all [xs] = - map single xs + fun pick_all [xs] = map single xs | pick_all (xs::xss) = - let val rec_pick = pick_all xss in - maps (fn x => map (cons x) rec_pick) xs - end - | pick_all _ = - raise REFUTE ("interpretation_apply", "empty list (in pick_all)") + let val rec_pick = pick_all xss in + maps (fn x => map (cons x) rec_pick) xs + end + | pick_all _ = raise REFUTE ("interpretation_apply", "empty list (in pick_all)") (* interpretation -> prop_formula list *) - fun interpretation_to_prop_formula_list (Leaf xs) = - xs + fun interpretation_to_prop_formula_list (Leaf xs) = xs | interpretation_to_prop_formula_list (Node trees) = - map PropLogic.all (pick_all - (map interpretation_to_prop_formula_list trees)) + map PropLogic.all (pick_all + (map interpretation_to_prop_formula_list trees)) in case i1 of Leaf _ => - raise REFUTE ("interpretation_apply", "first interpretation is a leaf") + raise REFUTE ("interpretation_apply", "first interpretation is a leaf") | Node xs => - prop_formula_list_dot_product_interpretation_list - (interpretation_to_prop_formula_list i2, xs) + prop_formula_list_dot_product_interpretation_list + (interpretation_to_prop_formula_list i2, xs) end; (* ------------------------------------------------------------------------- *) (* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions *) (* ------------------------------------------------------------------------- *) - (* Term.term -> int -> Term.term *) +(* Term.term -> int -> Term.term *) - fun eta_expand t i = +fun eta_expand t i = let val Ts = Term.binder_types (Term.fastype_of t) val t' = Term.incr_boundvars i t @@ -1605,1056 +1532,1036 @@ (* their arguments) of the size of the argument types *) (* ------------------------------------------------------------------------- *) - fun size_of_dtyp thy typ_sizes descr typ_assoc constructors = - Integer.sum (map (fn (_, dtyps) => - Integer.prod (map (size_of_type thy (typ_sizes, []) o - (typ_of_dtyp descr typ_assoc)) dtyps)) - constructors); +fun size_of_dtyp ctxt typ_sizes descr typ_assoc constructors = + Integer.sum (map (fn (_, dtyps) => + Integer.prod (map (size_of_type ctxt (typ_sizes, []) o + (typ_of_dtyp descr typ_assoc)) dtyps)) + constructors); (* ------------------------------------------------------------------------- *) (* INTERPRETERS: Actual Interpreters *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* simply typed lambda calculus: Isabelle's basic term syntax, with type *) +(* variables, function types, and propT *) - (* simply typed lambda calculus: Isabelle's basic term syntax, with type *) - (* variables, function types, and propT *) - - fun stlc_interpreter thy model args t = +fun stlc_interpreter ctxt model args t = let - val (typs, terms) = model + val thy = ProofContext.theory_of ctxt + val (typs, terms) = model val {maxvars, def_eq, next_idx, bounds, wellformed} = args (* Term.typ -> (interpretation * model * arguments) option *) fun interpret_groundterm T = - let - (* unit -> (interpretation * model * arguments) option *) - fun interpret_groundtype () = let - (* the model must specify a size for ground types *) - val size = if T = Term.propT then 2 - else the (AList.lookup (op =) typs T) - val next = next_idx+size - (* check if 'maxvars' is large enough *) - val _ = (if next-1>maxvars andalso maxvars>0 then - raise MAXVARS_EXCEEDED else ()) - (* prop_formula list *) - val fms = map BoolVar (next_idx upto (next_idx+size-1)) - (* interpretation *) - val intr = Leaf fms - (* prop_formula list -> prop_formula *) - fun one_of_two_false [] = True - | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => - SOr (SNot x, SNot x')) xs), one_of_two_false xs) - (* prop_formula *) - val wf = one_of_two_false fms + (* unit -> (interpretation * model * arguments) option *) + fun interpret_groundtype () = + let + (* the model must specify a size for ground types *) + val size = + if T = Term.propT then 2 + else the (AList.lookup (op =) typs T) + val next = next_idx + size + (* check if 'maxvars' is large enough *) + val _ = (if next - 1 > maxvars andalso maxvars > 0 then + raise MAXVARS_EXCEEDED else ()) + (* prop_formula list *) + val fms = map BoolVar (next_idx upto (next_idx + size - 1)) + (* interpretation *) + val intr = Leaf fms + (* prop_formula list -> prop_formula *) + fun one_of_two_false [] = True + | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => + SOr (SNot x, SNot x')) xs), one_of_two_false xs) + (* prop_formula *) + val wf = one_of_two_false fms + in + (* extend the model, increase 'next_idx', add well-formedness *) + (* condition *) + SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, + def_eq = def_eq, next_idx = next, bounds = bounds, + wellformed = SAnd (wellformed, wf)}) + end in - (* extend the model, increase 'next_idx', add well-formedness *) - (* condition *) - SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, - def_eq = def_eq, next_idx = next, bounds = bounds, - wellformed = SAnd (wellformed, wf)}) + case T of + Type ("fun", [T1, T2]) => + let + (* we create 'size_of_type ... T1' different copies of the *) + (* interpretation for 'T2', which are then combined into a single *) + (* new interpretation *) + (* make fresh copies, with different variable indices *) + (* 'idx': next variable index *) + (* 'n' : number of copies *) + (* int -> int -> (int * interpretation list * prop_formula *) + fun make_copies idx 0 = (idx, [], True) + | make_copies idx n = + let + val (copy, _, new_args) = interpret ctxt (typs, []) + {maxvars = maxvars, def_eq = false, next_idx = idx, + bounds = [], wellformed = True} (Free ("dummy", T2)) + val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1) + in + (idx', copy :: copies, SAnd (#wellformed new_args, wf')) + end + val (next, copies, wf) = make_copies next_idx + (size_of_type ctxt model T1) + (* combine copies into a single interpretation *) + val intr = Node copies + in + (* extend the model, increase 'next_idx', add well-formedness *) + (* condition *) + SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, + def_eq = def_eq, next_idx = next, bounds = bounds, + wellformed = SAnd (wellformed, wf)}) + end + | Type _ => interpret_groundtype () + | TFree _ => interpret_groundtype () + | TVar _ => interpret_groundtype () end - in - case T of - Type ("fun", [T1, T2]) => - let - (* we create 'size_of_type ... T1' different copies of the *) - (* interpretation for 'T2', which are then combined into a single *) - (* new interpretation *) - (* make fresh copies, with different variable indices *) - (* 'idx': next variable index *) - (* 'n' : number of copies *) - (* int -> int -> (int * interpretation list * prop_formula *) - fun make_copies idx 0 = - (idx, [], True) - | make_copies idx n = - let - val (copy, _, new_args) = interpret thy (typs, []) - {maxvars = maxvars, def_eq = false, next_idx = idx, - bounds = [], wellformed = True} (Free ("dummy", T2)) - val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1) - in - (idx', copy :: copies, SAnd (#wellformed new_args, wf')) - end - val (next, copies, wf) = make_copies next_idx - (size_of_type thy model T1) - (* combine copies into a single interpretation *) - val intr = Node copies - in - (* extend the model, increase 'next_idx', add well-formedness *) - (* condition *) - SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, - def_eq = def_eq, next_idx = next, bounds = bounds, - wellformed = SAnd (wellformed, wf)}) - end - | Type _ => interpret_groundtype () - | TFree _ => interpret_groundtype () - | TVar _ => interpret_groundtype () - end in case AList.lookup (op =) terms t of SOME intr => - (* return an existing interpretation *) - SOME (intr, model, args) + (* return an existing interpretation *) + SOME (intr, model, args) | NONE => - (case t of - Const (_, T) => - interpret_groundterm T - | Free (_, T) => - interpret_groundterm T - | Var (_, T) => - interpret_groundterm T - | Bound i => - SOME (List.nth (#bounds args, i), model, args) - | Abs (x, T, body) => - let - (* create all constants of type 'T' *) - val constants = make_constants thy model T - (* interpret the 'body' separately for each constant *) - val (bodies, (model', args')) = fold_map - (fn c => fn (m, a) => - let - (* add 'c' to 'bounds' *) - val (i', m', a') = interpret thy m {maxvars = #maxvars a, - def_eq = #def_eq a, next_idx = #next_idx a, - bounds = (c :: #bounds a), wellformed = #wellformed a} body - in - (* keep the new model m' and 'next_idx' and 'wellformed', *) - (* but use old 'bounds' *) - (i', (m', {maxvars = maxvars, def_eq = def_eq, - next_idx = #next_idx a', bounds = bounds, - wellformed = #wellformed a'})) - end) - constants (model, args) - in - SOME (Node bodies, model', args') - end - | t1 $ t2 => - let - (* interpret 't1' and 't2' separately *) - val (intr1, model1, args1) = interpret thy model args t1 - val (intr2, model2, args2) = interpret thy model1 args1 t2 - in - SOME (interpretation_apply (intr1, intr2), model2, args2) - end) + (case t of + Const (_, T) => interpret_groundterm T + | Free (_, T) => interpret_groundterm T + | Var (_, T) => interpret_groundterm T + | Bound i => SOME (List.nth (#bounds args, i), model, args) + | Abs (x, T, body) => + let + (* create all constants of type 'T' *) + val constants = make_constants ctxt model T + (* interpret the 'body' separately for each constant *) + val (bodies, (model', args')) = fold_map + (fn c => fn (m, a) => + let + (* add 'c' to 'bounds' *) + val (i', m', a') = interpret ctxt m {maxvars = #maxvars a, + def_eq = #def_eq a, next_idx = #next_idx a, + bounds = (c :: #bounds a), wellformed = #wellformed a} body + in + (* keep the new model m' and 'next_idx' and 'wellformed', *) + (* but use old 'bounds' *) + (i', (m', {maxvars = maxvars, def_eq = def_eq, + next_idx = #next_idx a', bounds = bounds, + wellformed = #wellformed a'})) + end) + constants (model, args) + in + SOME (Node bodies, model', args') + end + | t1 $ t2 => + let + (* interpret 't1' and 't2' separately *) + val (intr1, model1, args1) = interpret ctxt model args t1 + val (intr2, model2, args2) = interpret ctxt model1 args1 t2 + in + SOME (interpretation_apply (intr1, intr2), model2, args2) + end) end; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) - - fun Pure_interpreter thy model args t = - case t of - Const (@{const_name all}, _) $ t1 => +fun Pure_interpreter ctxt model args t = + case t of + Const (@{const_name all}, _) $ t1 => let - val (i, m, a) = interpret thy model args t1 + val (i, m, a) = interpret ctxt model args t1 in case i of Node xs => - (* 3-valued logic *) - let - val fmTrue = PropLogic.all (map toTrue xs) - val fmFalse = PropLogic.exists (map toFalse xs) - in - SOME (Leaf [fmTrue, fmFalse], m, a) - end + (* 3-valued logic *) + let + val fmTrue = PropLogic.all (map toTrue xs) + val fmFalse = PropLogic.exists (map toFalse xs) + in + SOME (Leaf [fmTrue, fmFalse], m, a) + end | _ => raise REFUTE ("Pure_interpreter", "\"all\" is followed by a non-function") end - | Const (@{const_name all}, _) => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name "=="}, _) $ t1 $ t2 => + | Const (@{const_name all}, _) => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name "=="}, _) $ t1 $ t2 => let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 in (* we use either 'make_def_equality' or 'make_equality' *) SOME ((if #def_eq args then make_def_equality else make_equality) (i1, i2), m2, a2) end - | Const (@{const_name "=="}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name "=="}, _) => - SOME (interpret thy model args (eta_expand t 2)) - | Const (@{const_name "==>"}, _) $ t1 $ t2 => + | Const (@{const_name "=="}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name "=="}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) + | Const (@{const_name "==>"}, _) $ t1 $ t2 => (* 3-valued logic *) let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2) val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2) in SOME (Leaf [fmTrue, fmFalse], m2, a2) end - | Const (@{const_name "==>"}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name "==>"}, _) => - SOME (interpret thy model args (eta_expand t 2)) - | _ => NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | Const (@{const_name "==>"}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name "==>"}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) + | _ => NONE; - fun HOLogic_interpreter thy model args t = - (* Providing interpretations directly is more efficient than unfolding the *) - (* logical constants. In HOL however, logical constants can themselves be *) - (* arguments. They are then translated using eta-expansion. *) - case t of - Const (@{const_name Trueprop}, _) => +fun HOLogic_interpreter ctxt model args t = +(* Providing interpretations directly is more efficient than unfolding the *) +(* logical constants. In HOL however, logical constants can themselves be *) +(* arguments. They are then translated using eta-expansion. *) + case t of + Const (@{const_name Trueprop}, _) => SOME (Node [TT, FF], model, args) - | Const (@{const_name Not}, _) => + | Const (@{const_name Not}, _) => SOME (Node [FF, TT], model, args) - (* redundant, since 'True' is also an IDT constructor *) - | Const (@{const_name True}, _) => + (* redundant, since 'True' is also an IDT constructor *) + | Const (@{const_name True}, _) => SOME (TT, model, args) - (* redundant, since 'False' is also an IDT constructor *) - | Const (@{const_name False}, _) => + (* redundant, since 'False' is also an IDT constructor *) + | Const (@{const_name False}, _) => SOME (FF, model, args) - | Const (@{const_name All}, _) $ t1 => (* similar to "all" (Pure) *) + | Const (@{const_name All}, _) $ t1 => (* similar to "all" (Pure) *) let - val (i, m, a) = interpret thy model args t1 + val (i, m, a) = interpret ctxt model args t1 in case i of Node xs => - (* 3-valued logic *) - let - val fmTrue = PropLogic.all (map toTrue xs) - val fmFalse = PropLogic.exists (map toFalse xs) - in - SOME (Leaf [fmTrue, fmFalse], m, a) - end + (* 3-valued logic *) + let + val fmTrue = PropLogic.all (map toTrue xs) + val fmFalse = PropLogic.exists (map toFalse xs) + in + SOME (Leaf [fmTrue, fmFalse], m, a) + end | _ => raise REFUTE ("HOLogic_interpreter", "\"All\" is followed by a non-function") end - | Const (@{const_name All}, _) => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name Ex}, _) $ t1 => + | Const (@{const_name All}, _) => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name Ex}, _) $ t1 => let - val (i, m, a) = interpret thy model args t1 + val (i, m, a) = interpret ctxt model args t1 in case i of Node xs => - (* 3-valued logic *) - let - val fmTrue = PropLogic.exists (map toTrue xs) - val fmFalse = PropLogic.all (map toFalse xs) - in - SOME (Leaf [fmTrue, fmFalse], m, a) - end + (* 3-valued logic *) + let + val fmTrue = PropLogic.exists (map toTrue xs) + val fmFalse = PropLogic.all (map toFalse xs) + in + SOME (Leaf [fmTrue, fmFalse], m, a) + end | _ => raise REFUTE ("HOLogic_interpreter", "\"Ex\" is followed by a non-function") end - | Const (@{const_name Ex}, _) => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name HOL.eq}, _) $ t1 $ t2 => (* similar to "==" (Pure) *) + | Const (@{const_name Ex}, _) => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name HOL.eq}, _) $ t1 $ t2 => (* similar to "==" (Pure) *) let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 in SOME (make_equality (i1, i2), m2, a2) end - | Const (@{const_name HOL.eq}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name HOL.eq}, _) => - SOME (interpret thy model args (eta_expand t 2)) - | Const (@{const_name HOL.conj}, _) $ t1 $ t2 => + | Const (@{const_name HOL.eq}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name HOL.eq}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) + | Const (@{const_name HOL.conj}, _) $ t1 $ t2 => (* 3-valued logic *) let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 val fmTrue = PropLogic.SAnd (toTrue i1, toTrue i2) val fmFalse = PropLogic.SOr (toFalse i1, toFalse i2) in SOME (Leaf [fmTrue, fmFalse], m2, a2) end - | Const (@{const_name HOL.conj}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name HOL.conj}, _) => - SOME (interpret thy model args (eta_expand t 2)) + | Const (@{const_name HOL.conj}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name HOL.conj}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) (* this would make "undef" propagate, even for formulae like *) (* "False & undef": *) (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *) - | Const (@{const_name HOL.disj}, _) $ t1 $ t2 => + | Const (@{const_name HOL.disj}, _) $ t1 $ t2 => (* 3-valued logic *) let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 val fmTrue = PropLogic.SOr (toTrue i1, toTrue i2) val fmFalse = PropLogic.SAnd (toFalse i1, toFalse i2) in SOME (Leaf [fmTrue, fmFalse], m2, a2) end - | Const (@{const_name HOL.disj}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name HOL.disj}, _) => - SOME (interpret thy model args (eta_expand t 2)) + | Const (@{const_name HOL.disj}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name HOL.disj}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) (* this would make "undef" propagate, even for formulae like *) (* "True | undef": *) (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *) - | Const (@{const_name HOL.implies}, _) $ t1 $ t2 => (* similar to "==>" (Pure) *) + | Const (@{const_name HOL.implies}, _) $ t1 $ t2 => (* similar to "==>" (Pure) *) (* 3-valued logic *) let - val (i1, m1, a1) = interpret thy model args t1 - val (i2, m2, a2) = interpret thy m1 a1 t2 + val (i1, m1, a1) = interpret ctxt model args t1 + val (i2, m2, a2) = interpret ctxt m1 a1 t2 val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2) val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2) in SOME (Leaf [fmTrue, fmFalse], m2, a2) end - | Const (@{const_name HOL.implies}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name HOL.implies}, _) => - SOME (interpret thy model args (eta_expand t 2)) + | Const (@{const_name HOL.implies}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name HOL.implies}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) (* this would make "undef" propagate, even for formulae like *) (* "False --> undef": *) (* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *) - | _ => NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* interprets variables and constants whose type is an IDT (this is *) - (* relatively easy and merely requires us to compute the size of the IDT); *) - (* constructors of IDTs however are properly interpreted by *) - (* 'IDT_constructor_interpreter' *) +(* interprets variables and constants whose type is an IDT (this is *) +(* relatively easy and merely requires us to compute the size of the IDT); *) +(* constructors of IDTs however are properly interpreted by *) +(* 'IDT_constructor_interpreter' *) - fun IDT_interpreter thy model args t = +fun IDT_interpreter ctxt model args t = let + val thy = ProofContext.theory_of ctxt val (typs, terms) = model (* Term.typ -> (interpretation * model * arguments) option *) fun interpret_term (Type (s, Ts)) = - (case Datatype.get_info thy s of - SOME info => (* inductive datatype *) - let - (* int option -- only recursive IDTs have an associated depth *) - val depth = AList.lookup (op =) typs (Type (s, Ts)) - (* sanity check: depth must be at least 0 *) - val _ = (case depth of SOME n => - if n<0 then - raise REFUTE ("IDT_interpreter", "negative depth") - else () - | _ => ()) - in - (* termination condition to avoid infinite recursion *) - if depth = (SOME 0) then - (* return a leaf of size 0 *) - SOME (Leaf [], model, args) - else - let - val index = #index info - val descr = #descr info - val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) - val typ_assoc = dtyps ~~ Ts - (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps - then - raise REFUTE ("IDT_interpreter", - "datatype argument (for type " - ^ Syntax.string_of_typ_global thy (Type (s, Ts)) - ^ ") is not a variable") - else () - (* if the model specifies a depth for the current type, *) - (* decrement it to avoid infinite recursion *) - val typs' = case depth of NONE => typs | SOME n => - AList.update (op =) (Type (s, Ts), n-1) typs - (* recursively compute the size of the datatype *) - val size = size_of_dtyp thy typs' descr typ_assoc constrs - val next_idx = #next_idx args - val next = next_idx+size - (* check if 'maxvars' is large enough *) - val _ = (if next-1 > #maxvars args andalso - #maxvars args > 0 then raise MAXVARS_EXCEEDED else ()) - (* prop_formula list *) - val fms = map BoolVar (next_idx upto (next_idx+size-1)) - (* interpretation *) - val intr = Leaf fms - (* prop_formula list -> prop_formula *) - fun one_of_two_false [] = True - | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => - SOr (SNot x, SNot x')) xs), one_of_two_false xs) - (* prop_formula *) - val wf = one_of_two_false fms - in - (* extend the model, increase 'next_idx', add well-formedness *) - (* condition *) - SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args, - def_eq = #def_eq args, next_idx = next, bounds = #bounds args, - wellformed = SAnd (#wellformed args, wf)}) - end - end - | NONE => (* not an inductive datatype *) - NONE) + (case Datatype.get_info thy s of + SOME info => (* inductive datatype *) + let + (* int option -- only recursive IDTs have an associated depth *) + val depth = AList.lookup (op =) typs (Type (s, Ts)) + (* sanity check: depth must be at least 0 *) + val _ = + (case depth of SOME n => + if n < 0 then + raise REFUTE ("IDT_interpreter", "negative depth") + else () + | _ => ()) + in + (* termination condition to avoid infinite recursion *) + if depth = (SOME 0) then + (* return a leaf of size 0 *) + SOME (Leaf [], model, args) + else + let + val index = #index info + val descr = #descr info + val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) + val typ_assoc = dtyps ~~ Ts + (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) + val _ = + if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps + then + raise REFUTE ("IDT_interpreter", + "datatype argument (for type " + ^ Syntax.string_of_typ ctxt (Type (s, Ts)) + ^ ") is not a variable") + else () + (* if the model specifies a depth for the current type, *) + (* decrement it to avoid infinite recursion *) + val typs' = case depth of NONE => typs | SOME n => + AList.update (op =) (Type (s, Ts), n-1) typs + (* recursively compute the size of the datatype *) + val size = size_of_dtyp ctxt typs' descr typ_assoc constrs + val next_idx = #next_idx args + val next = next_idx+size + (* check if 'maxvars' is large enough *) + val _ = (if next-1 > #maxvars args andalso + #maxvars args > 0 then raise MAXVARS_EXCEEDED else ()) + (* prop_formula list *) + val fms = map BoolVar (next_idx upto (next_idx+size-1)) + (* interpretation *) + val intr = Leaf fms + (* prop_formula list -> prop_formula *) + fun one_of_two_false [] = True + | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => + SOr (SNot x, SNot x')) xs), one_of_two_false xs) + (* prop_formula *) + val wf = one_of_two_false fms + in + (* extend the model, increase 'next_idx', add well-formedness *) + (* condition *) + SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args, + def_eq = #def_eq args, next_idx = next, bounds = #bounds args, + wellformed = SAnd (#wellformed args, wf)}) + end + end + | NONE => (* not an inductive datatype *) + NONE) | interpret_term _ = (* a (free or schematic) type variable *) - NONE + NONE in case AList.lookup (op =) terms t of SOME intr => - (* return an existing interpretation *) - SOME (intr, model, args) + (* return an existing interpretation *) + SOME (intr, model, args) | NONE => - (case t of - Free (_, T) => interpret_term T - | Var (_, T) => interpret_term T - | Const (_, T) => interpret_term T - | _ => NONE) + (case t of + Free (_, T) => interpret_term T + | Var (_, T) => interpret_term T + | Const (_, T) => interpret_term T + | _ => NONE) end; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* This function imposes an order on the elements of a datatype fragment *) +(* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or *) +(* (x_1, ..., x_n) < (y_1, ..., y_m). With this order, a constructor is *) +(* a function C_i that maps some argument indices x_1, ..., x_n to the *) +(* datatype element given by index C_i x_1 ... x_n. The idea remains the *) +(* same for recursive datatypes, although the computation of indices gets *) +(* a little tricky. *) - (* This function imposes an order on the elements of a datatype fragment *) - (* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or *) - (* (x_1, ..., x_n) < (y_1, ..., y_m). With this order, a constructor is *) - (* a function C_i that maps some argument indices x_1, ..., x_n to the *) - (* datatype element given by index C_i x_1 ... x_n. The idea remains the *) - (* same for recursive datatypes, although the computation of indices gets *) - (* a little tricky. *) - - fun IDT_constructor_interpreter thy model args t = +fun IDT_constructor_interpreter ctxt model args t = let + val thy = ProofContext.theory_of ctxt (* returns a list of canonical representations for terms of the type 'T' *) (* It would be nice if we could just use 'print' for this, but 'print' *) (* for IDTs calls 'IDT_constructor_interpreter' again, and this could *) (* lead to infinite recursion when we have (mutually) recursive IDTs. *) (* (Term.typ * int) list -> Term.typ -> Term.term list *) fun canonical_terms typs T = - (case T of - Type ("fun", [T1, T2]) => - (* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *) - (* least not for 'T2' *) - let - (* returns a list of lists, each one consisting of n (possibly *) - (* identical) elements from 'xs' *) - (* int -> 'a list -> 'a list list *) - fun pick_all 1 xs = - map single xs - | pick_all n xs = - let val rec_pick = pick_all (n-1) xs in - maps (fn x => map (cons x) rec_pick) xs - end - (* ["x1", ..., "xn"] *) - val terms1 = canonical_terms typs T1 - (* ["y1", ..., "ym"] *) - val terms2 = canonical_terms typs T2 - (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *) - (* [("x1", "ym"), ..., ("xn", "ym")]] *) - val functions = map (curry (op ~~) terms1) - (pick_all (length terms1) terms2) - (* [["(x1, y1)", ..., "(xn, y1)"], ..., *) - (* ["(x1, ym)", ..., "(xn, ym)"]] *) - val pairss = map (map HOLogic.mk_prod) functions - (* Term.typ *) - val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) - val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT - (* Term.term *) - val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT) - val HOLogic_insert = - Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) - in - (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *) - map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps - HOLogic_empty_set) pairss - end - | Type (s, Ts) => - (case Datatype.get_info thy s of - SOME info => - (case AList.lookup (op =) typs T of - SOME 0 => - (* termination condition to avoid infinite recursion *) - [] (* at depth 0, every IDT is empty *) - | _ => + (case T of + Type ("fun", [T1, T2]) => + (* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *) + (* least not for 'T2' *) let - val index = #index info - val descr = #descr info - val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) - val typ_assoc = dtyps ~~ Ts - (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps - then - raise REFUTE ("IDT_constructor_interpreter", - "datatype argument (for type " - ^ Syntax.string_of_typ_global thy T - ^ ") is not a variable") - else () - (* decrement depth for the IDT 'T' *) - val typs' = (case AList.lookup (op =) typs T of NONE => typs - | SOME n => AList.update (op =) (T, n-1) typs) - fun constructor_terms terms [] = terms - | constructor_terms terms (d::ds) = + (* returns a list of lists, each one consisting of n (possibly *) + (* identical) elements from 'xs' *) + (* int -> 'a list -> 'a list list *) + fun pick_all 1 xs = map single xs + | pick_all n xs = + let val rec_pick = pick_all (n-1) xs in + maps (fn x => map (cons x) rec_pick) xs + end + (* ["x1", ..., "xn"] *) + val terms1 = canonical_terms typs T1 + (* ["y1", ..., "ym"] *) + val terms2 = canonical_terms typs T2 + (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *) + (* [("x1", "ym"), ..., ("xn", "ym")]] *) + val functions = map (curry (op ~~) terms1) + (pick_all (length terms1) terms2) + (* [["(x1, y1)", ..., "(xn, y1)"], ..., *) + (* ["(x1, ym)", ..., "(xn, ym)"]] *) + val pairss = map (map HOLogic.mk_prod) functions + (* Term.typ *) + val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) + val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT + (* Term.term *) + val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT) + val HOLogic_insert = + Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) + in + (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *) + map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps + HOLogic_empty_set) pairss + end + | Type (s, Ts) => + (case Datatype.get_info thy s of + SOME info => + (case AList.lookup (op =) typs T of + SOME 0 => + (* termination condition to avoid infinite recursion *) + [] (* at depth 0, every IDT is empty *) + | _ => let - val dT = typ_of_dtyp descr typ_assoc d - val d_terms = canonical_terms typs' dT + val index = #index info + val descr = #descr info + val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) + val typ_assoc = dtyps ~~ Ts + (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) + val _ = + if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps + then + raise REFUTE ("IDT_constructor_interpreter", + "datatype argument (for type " + ^ Syntax.string_of_typ ctxt T + ^ ") is not a variable") + else () + (* decrement depth for the IDT 'T' *) + val typs' = + (case AList.lookup (op =) typs T of NONE => typs + | SOME n => AList.update (op =) (T, n-1) typs) + fun constructor_terms terms [] = terms + | constructor_terms terms (d::ds) = + let + val dT = typ_of_dtyp descr typ_assoc d + val d_terms = canonical_terms typs' dT + in + (* C_i x_1 ... x_n < C_i y_1 ... y_n if *) + (* (x_1, ..., x_n) < (y_1, ..., y_n) *) + constructor_terms + (map_product (curry op $) terms d_terms) ds + end in - (* C_i x_1 ... x_n < C_i y_1 ... y_n if *) - (* (x_1, ..., x_n) < (y_1, ..., y_n) *) - constructor_terms - (map_product (curry op $) terms d_terms) ds - end - in - (* C_i ... < C_j ... if i < j *) - maps (fn (cname, ctyps) => - let - val cTerm = Const (cname, - map (typ_of_dtyp descr typ_assoc) ctyps ---> T) - in - constructor_terms [cTerm] ctyps - end) constrs - end) - | NONE => - (* not an inductive datatype; in this case the argument types in *) - (* 'Ts' may not be IDTs either, so 'print' should be safe *) - map (fn intr => print thy (typs, []) T intr (K false)) - (make_constants thy (typs, []) T)) + (* C_i ... < C_j ... if i < j *) + maps (fn (cname, ctyps) => + let + val cTerm = Const (cname, + map (typ_of_dtyp descr typ_assoc) ctyps ---> T) + in + constructor_terms [cTerm] ctyps + end) constrs + end) + | NONE => + (* not an inductive datatype; in this case the argument types in *) + (* 'Ts' may not be IDTs either, so 'print' should be safe *) + map (fn intr => print ctxt (typs, []) T intr (K false)) + (make_constants ctxt (typs, []) T)) | _ => (* TFree ..., TVar ... *) - map (fn intr => print thy (typs, []) T intr (K false)) - (make_constants thy (typs, []) T)) + map (fn intr => print ctxt (typs, []) T intr (K false)) + (make_constants ctxt (typs, []) T)) val (typs, terms) = model in case AList.lookup (op =) terms t of SOME intr => - (* return an existing interpretation *) - SOME (intr, model, args) + (* return an existing interpretation *) + SOME (intr, model, args) | NONE => - (case t of - Const (s, T) => - (case body_type T of - Type (s', Ts') => - (case Datatype.get_info thy s' of - SOME info => (* body type is an inductive datatype *) - let - val index = #index info - val descr = #descr info - val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) - val typ_assoc = dtyps ~~ Ts' - (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps - then - raise REFUTE ("IDT_constructor_interpreter", - "datatype argument (for type " - ^ Syntax.string_of_typ_global thy (Type (s', Ts')) - ^ ") is not a variable") - else () - (* split the constructors into those occuring before/after *) - (* 'Const (s, T)' *) - val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) => - not (cname = s andalso Sign.typ_instance thy (T, - map (typ_of_dtyp descr typ_assoc) ctypes - ---> Type (s', Ts')))) constrs - in - case constrs2 of - [] => - (* 'Const (s, T)' is not a constructor of this datatype *) - NONE - | (_, ctypes)::cs => - let - (* int option -- only /recursive/ IDTs have an associated *) - (* depth *) - val depth = AList.lookup (op =) typs (Type (s', Ts')) - (* this should never happen: at depth 0, this IDT fragment *) - (* is definitely empty, and in this case we don't need to *) - (* interpret its constructors *) - val _ = (case depth of SOME 0 => - raise REFUTE ("IDT_constructor_interpreter", - "depth is 0") - | _ => ()) - val typs' = (case depth of NONE => typs | SOME n => - AList.update (op =) (Type (s', Ts'), n-1) typs) - (* elements of the datatype come before elements generated *) - (* by 'Const (s, T)' iff they are generated by a *) - (* constructor in constrs1 *) - val offset = size_of_dtyp thy typs' descr typ_assoc constrs1 - (* compute the total (current) size of the datatype *) - val total = offset + - size_of_dtyp thy typs' descr typ_assoc constrs2 - (* sanity check *) - val _ = if total <> size_of_type thy (typs, []) - (Type (s', Ts')) then - raise REFUTE ("IDT_constructor_interpreter", - "total is not equal to current size") - else () - (* returns an interpretation where everything is mapped to *) - (* an "undefined" element of the datatype *) - fun make_undef [] = - Leaf (replicate total False) - | make_undef (d::ds) = + (case t of + Const (s, T) => + (case body_type T of + Type (s', Ts') => + (case Datatype.get_info thy s' of + SOME info => (* body type is an inductive datatype *) let - (* compute the current size of the type 'd' *) - val dT = typ_of_dtyp descr typ_assoc d - val size = size_of_type thy (typs, []) dT - in - Node (replicate size (make_undef ds)) - end - (* returns the interpretation for a constructor *) - fun make_constr [] offset = - if offset < total then - (Leaf (replicate offset False @ True :: - (replicate (total - offset - 1) False)), offset + 1) - else - raise REFUTE ("IDT_constructor_interpreter", - "offset >= total") - | make_constr (d::ds) offset = - let - (* Term.typ *) - val dT = typ_of_dtyp descr typ_assoc d - (* compute canonical term representations for all *) - (* elements of the type 'd' (with the reduced depth *) - (* for the IDT) *) - val terms' = canonical_terms typs' dT - (* sanity check *) - val _ = - if length terms' <> size_of_type thy (typs', []) dT + val index = #index info + val descr = #descr info + val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) + val typ_assoc = dtyps ~~ Ts' + (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) + val _ = if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_constructor_interpreter", - "length of terms' is not equal to old size") - else () - (* compute canonical term representations for all *) - (* elements of the type 'd' (with the current depth *) - (* for the IDT) *) - val terms = canonical_terms typs dT - (* sanity check *) - val _ = - if length terms <> size_of_type thy (typs, []) dT - then - raise REFUTE ("IDT_constructor_interpreter", - "length of terms is not equal to current size") - else () - (* sanity check *) - val _ = - if length terms < length terms' then - raise REFUTE ("IDT_constructor_interpreter", - "current size is less than old size") + "datatype argument (for type " + ^ Syntax.string_of_typ ctxt (Type (s', Ts')) + ^ ") is not a variable") else () - (* sanity check: every element of terms' must also be *) - (* present in terms *) - val _ = - if forall (member (op =) terms) terms' then () - else - raise REFUTE ("IDT_constructor_interpreter", - "element has disappeared") - (* sanity check: the order on elements of terms' is *) - (* the same in terms, for those elements *) - val _ = - let - fun search (x::xs) (y::ys) = - if x = y then search xs ys else search (x::xs) ys - | search (x::xs) [] = + (* split the constructors into those occuring before/after *) + (* 'Const (s, T)' *) + val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) => + not (cname = s andalso Sign.typ_instance thy (T, + map (typ_of_dtyp descr typ_assoc) ctypes + ---> Type (s', Ts')))) constrs + in + case constrs2 of + [] => + (* 'Const (s, T)' is not a constructor of this datatype *) + NONE + | (_, ctypes)::cs => + let + (* int option -- only /recursive/ IDTs have an associated *) + (* depth *) + val depth = AList.lookup (op =) typs (Type (s', Ts')) + (* this should never happen: at depth 0, this IDT fragment *) + (* is definitely empty, and in this case we don't need to *) + (* interpret its constructors *) + val _ = (case depth of SOME 0 => + raise REFUTE ("IDT_constructor_interpreter", + "depth is 0") + | _ => ()) + val typs' = (case depth of NONE => typs | SOME n => + AList.update (op =) (Type (s', Ts'), n-1) typs) + (* elements of the datatype come before elements generated *) + (* by 'Const (s, T)' iff they are generated by a *) + (* constructor in constrs1 *) + val offset = size_of_dtyp ctxt typs' descr typ_assoc constrs1 + (* compute the total (current) size of the datatype *) + val total = offset + + size_of_dtyp ctxt typs' descr typ_assoc constrs2 + (* sanity check *) + val _ = if total <> size_of_type ctxt (typs, []) + (Type (s', Ts')) then raise REFUTE ("IDT_constructor_interpreter", - "element order not preserved") - | search [] _ = () - in search terms' terms end - (* int * interpretation list *) - val (intrs, new_offset) = - fold_map (fn t_elem => fn off => - (* if 't_elem' existed at the previous depth, *) - (* proceed recursively, otherwise map the entire *) - (* subtree to "undefined" *) - if member (op =) terms' t_elem then - make_constr ds off - else - (make_undef ds, off)) - terms offset - in - (Node intrs, new_offset) + "total is not equal to current size") + else () + (* returns an interpretation where everything is mapped to *) + (* an "undefined" element of the datatype *) + fun make_undef [] = Leaf (replicate total False) + | make_undef (d::ds) = + let + (* compute the current size of the type 'd' *) + val dT = typ_of_dtyp descr typ_assoc d + val size = size_of_type ctxt (typs, []) dT + in + Node (replicate size (make_undef ds)) + end + (* returns the interpretation for a constructor *) + fun make_constr [] offset = + if offset < total then + (Leaf (replicate offset False @ True :: + (replicate (total - offset - 1) False)), offset + 1) + else + raise REFUTE ("IDT_constructor_interpreter", + "offset >= total") + | make_constr (d::ds) offset = + let + (* Term.typ *) + val dT = typ_of_dtyp descr typ_assoc d + (* compute canonical term representations for all *) + (* elements of the type 'd' (with the reduced depth *) + (* for the IDT) *) + val terms' = canonical_terms typs' dT + (* sanity check *) + val _ = + if length terms' <> size_of_type ctxt (typs', []) dT + then + raise REFUTE ("IDT_constructor_interpreter", + "length of terms' is not equal to old size") + else () + (* compute canonical term representations for all *) + (* elements of the type 'd' (with the current depth *) + (* for the IDT) *) + val terms = canonical_terms typs dT + (* sanity check *) + val _ = + if length terms <> size_of_type ctxt (typs, []) dT + then + raise REFUTE ("IDT_constructor_interpreter", + "length of terms is not equal to current size") + else () + (* sanity check *) + val _ = + if length terms < length terms' then + raise REFUTE ("IDT_constructor_interpreter", + "current size is less than old size") + else () + (* sanity check: every element of terms' must also be *) + (* present in terms *) + val _ = + if forall (member (op =) terms) terms' then () + else + raise REFUTE ("IDT_constructor_interpreter", + "element has disappeared") + (* sanity check: the order on elements of terms' is *) + (* the same in terms, for those elements *) + val _ = + let + fun search (x::xs) (y::ys) = + if x = y then search xs ys else search (x::xs) ys + | search (x::xs) [] = + raise REFUTE ("IDT_constructor_interpreter", + "element order not preserved") + | search [] _ = () + in search terms' terms end + (* int * interpretation list *) + val (intrs, new_offset) = + fold_map (fn t_elem => fn off => + (* if 't_elem' existed at the previous depth, *) + (* proceed recursively, otherwise map the entire *) + (* subtree to "undefined" *) + if member (op =) terms' t_elem then + make_constr ds off + else + (make_undef ds, off)) + terms offset + in + (Node intrs, new_offset) + end + in + SOME (fst (make_constr ctypes offset), model, args) + end end - in - SOME (fst (make_constr ctypes offset), model, args) - end - end - | NONE => (* body type is not an inductive datatype *) - NONE) - | _ => (* body type is a (free or schematic) type variable *) + | NONE => (* body type is not an inductive datatype *) + NONE) + | _ => (* body type is a (free or schematic) type variable *) + NONE) + | _ => (* term is not a constant *) NONE) - | _ => (* term is not a constant *) - NONE) end; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* Difficult code ahead. Make sure you understand the *) +(* 'IDT_constructor_interpreter' and the order in which it enumerates *) +(* elements of an IDT before you try to understand this function. *) - (* Difficult code ahead. Make sure you understand the *) - (* 'IDT_constructor_interpreter' and the order in which it enumerates *) - (* elements of an IDT before you try to understand this function. *) - - fun IDT_recursion_interpreter thy model args t = +fun IDT_recursion_interpreter ctxt model args t = + let + val thy = ProofContext.theory_of ctxt + in (* careful: here we descend arbitrarily deep into 't', possibly before *) (* any other interpreter for atomic terms has had a chance to look at *) (* 't' *) case strip_comb t of (Const (s, T), params) => - (* iterate over all datatypes in 'thy' *) - Symtab.fold (fn (_, info) => fn result => - case result of - SOME _ => - result (* just keep 'result' *) - | NONE => - if member (op =) (#rec_names info) s then - (* we do have a recursion operator of one of the (mutually *) - (* recursive) datatypes given by 'info' *) - let - (* number of all constructors, including those of different *) - (* (mutually recursive) datatypes within the same descriptor *) - val mconstrs_count = - Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info)) - in - if mconstrs_count < length params then - (* too many actual parameters; for now we'll use the *) - (* 'stlc_interpreter' to strip off one application *) - NONE - else if mconstrs_count > length params then - (* too few actual parameters; we use eta expansion *) - (* Note that the resulting expansion of lambda abstractions *) - (* by the 'stlc_interpreter' may be rather slow (depending *) - (* on the argument types and the size of the IDT, of *) - (* course). *) - SOME (interpret thy model args (eta_expand t - (mconstrs_count - length params))) - else (* mconstrs_count = length params *) + (* iterate over all datatypes in 'thy' *) + Symtab.fold (fn (_, info) => fn result => + case result of + SOME _ => + result (* just keep 'result' *) + | NONE => + if member (op =) (#rec_names info) s then + (* we do have a recursion operator of one of the (mutually *) + (* recursive) datatypes given by 'info' *) let - (* interpret each parameter separately *) - val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) => + (* number of all constructors, including those of different *) + (* (mutually recursive) datatypes within the same descriptor *) + val mconstrs_count = + Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info)) + in + if mconstrs_count < length params then + (* too many actual parameters; for now we'll use the *) + (* 'stlc_interpreter' to strip off one application *) + NONE + else if mconstrs_count > length params then + (* too few actual parameters; we use eta expansion *) + (* Note that the resulting expansion of lambda abstractions *) + (* by the 'stlc_interpreter' may be rather slow (depending *) + (* on the argument types and the size of the IDT, of *) + (* course). *) + SOME (interpret ctxt model args (eta_expand t + (mconstrs_count - length params))) + else (* mconstrs_count = length params *) let - val (i, m', a') = interpret thy m a p - in - (i, (m', a')) - end) params (model, args) - val (typs, _) = model' - (* 'index' is /not/ necessarily the index of the IDT that *) - (* the recursion operator is associated with, but merely *) - (* the index of some mutually recursive IDT *) - val index = #index info - val descr = #descr info - val (_, dtyps, _) = the (AList.lookup (op =) descr index) - (* sanity check: we assume that the order of constructors *) - (* in 'descr' is the same as the order of *) - (* corresponding parameters, otherwise the *) - (* association code below won't match the *) - (* right constructors/parameters; we also *) - (* assume that the order of recursion *) - (* operators in '#rec_names info' is the *) - (* same as the order of corresponding *) - (* datatypes in 'descr' *) - val _ = if map fst descr <> (0 upto (length descr - 1)) then - raise REFUTE ("IDT_recursion_interpreter", - "order of constructors and corresponding parameters/" ^ - "recursion operators and corresponding datatypes " ^ - "different?") - else () - (* sanity check: every element in 'dtyps' must be a *) - (* 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false - | _ => true) dtyps - then - raise REFUTE ("IDT_recursion_interpreter", - "datatype argument is not a variable") - else () - (* the type of a recursion operator is *) - (* [T1, ..., Tn, IDT] ---> Tresult *) - val IDT = List.nth (binder_types T, mconstrs_count) - (* by our assumption on the order of recursion operators *) - (* and datatypes, this is the index of the datatype *) - (* corresponding to the given recursion operator *) - val idt_index = find_index (fn s' => s' = s) (#rec_names info) - (* mutually recursive types must have the same type *) - (* parameters, unless the mutual recursion comes from *) - (* indirect recursion *) - fun rec_typ_assoc acc [] = - acc - | rec_typ_assoc acc ((d, T)::xs) = - (case AList.lookup op= acc d of - NONE => - (case d of - Datatype_Aux.DtTFree _ => - (* add the association, proceed *) - rec_typ_assoc ((d, T)::acc) xs - | Datatype_Aux.DtType (s, ds) => + (* interpret each parameter separately *) + val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) => + let + val (i, m', a') = interpret ctxt m a p + in + (i, (m', a')) + end) params (model, args) + val (typs, _) = model' + (* 'index' is /not/ necessarily the index of the IDT that *) + (* the recursion operator is associated with, but merely *) + (* the index of some mutually recursive IDT *) + val index = #index info + val descr = #descr info + val (_, dtyps, _) = the (AList.lookup (op =) descr index) + (* sanity check: we assume that the order of constructors *) + (* in 'descr' is the same as the order of *) + (* corresponding parameters, otherwise the *) + (* association code below won't match the *) + (* right constructors/parameters; we also *) + (* assume that the order of recursion *) + (* operators in '#rec_names info' is the *) + (* same as the order of corresponding *) + (* datatypes in 'descr' *) + val _ = if map fst descr <> (0 upto (length descr - 1)) then + raise REFUTE ("IDT_recursion_interpreter", + "order of constructors and corresponding parameters/" ^ + "recursion operators and corresponding datatypes " ^ + "different?") + else () + (* sanity check: every element in 'dtyps' must be a *) + (* 'DtTFree' *) + val _ = + if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false + | _ => true) dtyps + then + raise REFUTE ("IDT_recursion_interpreter", + "datatype argument is not a variable") + else () + (* the type of a recursion operator is *) + (* [T1, ..., Tn, IDT] ---> Tresult *) + val IDT = List.nth (binder_types T, mconstrs_count) + (* by our assumption on the order of recursion operators *) + (* and datatypes, this is the index of the datatype *) + (* corresponding to the given recursion operator *) + val idt_index = find_index (fn s' => s' = s) (#rec_names info) + (* mutually recursive types must have the same type *) + (* parameters, unless the mutual recursion comes from *) + (* indirect recursion *) + fun rec_typ_assoc acc [] = acc + | rec_typ_assoc acc ((d, T)::xs) = + (case AList.lookup op= acc d of + NONE => + (case d of + Datatype_Aux.DtTFree _ => + (* add the association, proceed *) + rec_typ_assoc ((d, T)::acc) xs + | Datatype_Aux.DtType (s, ds) => + let + val (s', Ts) = dest_Type T + in + if s=s' then + rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) + else + raise REFUTE ("IDT_recursion_interpreter", + "DtType/Type mismatch") + end + | Datatype_Aux.DtRec i => + let + val (_, ds, _) = the (AList.lookup (op =) descr i) + val (_, Ts) = dest_Type T + in + rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) + end) + | SOME T' => + if T=T' then + (* ignore the association since it's already *) + (* present, proceed *) + rec_typ_assoc acc xs + else + raise REFUTE ("IDT_recursion_interpreter", + "different type associations for the same dtyp")) + val typ_assoc = filter + (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false) + (rec_typ_assoc [] + (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT)) + (* sanity check: typ_assoc must associate types to the *) + (* elements of 'dtyps' (and only to those) *) + val _ = + if not (eq_set (op =) (dtyps, map fst typ_assoc)) + then + raise REFUTE ("IDT_recursion_interpreter", + "type association has extra/missing elements") + else () + (* interpret each constructor in the descriptor (including *) + (* those of mutually recursive datatypes) *) + (* (int * interpretation list) list *) + val mc_intrs = map (fn (idx, (_, _, cs)) => let - val (s', Ts) = dest_Type T + val c_return_typ = typ_of_dtyp descr typ_assoc + (Datatype_Aux.DtRec idx) in - if s=s' then - rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) - else + (idx, map (fn (cname, cargs) => + (#1 o interpret ctxt (typs, []) {maxvars=0, + def_eq=false, next_idx=1, bounds=[], + wellformed=True}) (Const (cname, map (typ_of_dtyp + descr typ_assoc) cargs ---> c_return_typ))) cs) + end) descr + (* associate constructors with corresponding parameters *) + (* (int * (interpretation * interpretation) list) list *) + val (mc_p_intrs, p_intrs') = fold_map + (fn (idx, c_intrs) => fn p_intrs' => + let + val len = length c_intrs + in + ((idx, c_intrs ~~ List.take (p_intrs', len)), + List.drop (p_intrs', len)) + end) mc_intrs p_intrs + (* sanity check: no 'p_intr' may be left afterwards *) + val _ = + if p_intrs' <> [] then + raise REFUTE ("IDT_recursion_interpreter", + "more parameter than constructor interpretations") + else () + (* The recursion operator, applied to 'mconstrs_count' *) + (* arguments, is a function that maps every element of the *) + (* inductive datatype to an element of some result type. *) + (* Recursion operators for mutually recursive IDTs are *) + (* translated simultaneously. *) + (* Since the order on datatype elements is given by an *) + (* order on constructors (and then by the order on *) + (* argument tuples), we can simply copy corresponding *) + (* subtrees from 'p_intrs', in the order in which they are *) + (* given. *) + (* interpretation * interpretation -> interpretation list *) + fun ci_pi (Leaf xs, pi) = + (* if the constructor does not match the arguments to a *) + (* defined element of the IDT, the corresponding value *) + (* of the parameter must be ignored *) + if List.exists (equal True) xs then [pi] else [] + | ci_pi (Node xs, Node ys) = maps ci_pi (xs ~~ ys) + | ci_pi (Node _, Leaf _) = raise REFUTE ("IDT_recursion_interpreter", - "DtType/Type mismatch") - end - | Datatype_Aux.DtRec i => + "constructor takes more arguments than the " ^ + "associated parameter") + (* (int * interpretation list) list *) + val rec_operators = map (fn (idx, c_p_intrs) => + (idx, maps ci_pi c_p_intrs)) mc_p_intrs + (* sanity check: every recursion operator must provide as *) + (* many values as the corresponding datatype *) + (* has elements *) + val _ = map (fn (idx, intrs) => let - val (_, ds, _) = the (AList.lookup (op =) descr i) - val (_, Ts) = dest_Type T + val T = typ_of_dtyp descr typ_assoc + (Datatype_Aux.DtRec idx) in - rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs) - end) - | SOME T' => - if T=T' then - (* ignore the association since it's already *) - (* present, proceed *) - rec_typ_assoc acc xs - else - raise REFUTE ("IDT_recursion_interpreter", - "different type associations for the same dtyp")) - val typ_assoc = filter - (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false) - (rec_typ_assoc [] - (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT)) - (* sanity check: typ_assoc must associate types to the *) - (* elements of 'dtyps' (and only to those) *) - val _ = if not (eq_set (op =) (dtyps, map fst typ_assoc)) - then - raise REFUTE ("IDT_recursion_interpreter", - "type association has extra/missing elements") - else () - (* interpret each constructor in the descriptor (including *) - (* those of mutually recursive datatypes) *) - (* (int * interpretation list) list *) - val mc_intrs = map (fn (idx, (_, _, cs)) => - let - val c_return_typ = typ_of_dtyp descr typ_assoc - (Datatype_Aux.DtRec idx) + if length intrs <> size_of_type ctxt (typs, []) T then + raise REFUTE ("IDT_recursion_interpreter", + "wrong number of interpretations for rec. operator") + else () + end) rec_operators + (* For non-recursive datatypes, we are pretty much done at *) + (* this point. For recursive datatypes however, we still *) + (* need to apply the interpretations in 'rec_operators' to *) + (* (recursively obtained) interpretations for recursive *) + (* constructor arguments. To do so more efficiently, we *) + (* copy 'rec_operators' into arrays first. Each Boolean *) + (* indicates whether the recursive arguments have been *) + (* considered already. *) + (* (int * (bool * interpretation) Array.array) list *) + val REC_OPERATORS = map (fn (idx, intrs) => + (idx, Array.fromList (map (pair false) intrs))) + rec_operators + (* takes an interpretation, and if some leaf of this *) + (* interpretation is the 'elem'-th element of the type, *) + (* the indices of the arguments leading to this leaf are *) + (* returned *) + (* interpretation -> int -> int list option *) + fun get_args (Leaf xs) elem = + if find_index (fn x => x = True) xs = elem then + SOME [] + else + NONE + | get_args (Node xs) elem = + let + (* interpretation list * int -> int list option *) + fun search ([], _) = + NONE + | search (x::xs, n) = + (case get_args x elem of + SOME result => SOME (n::result) + | NONE => search (xs, n+1)) + in + search (xs, 0) + end + (* returns the index of the constructor and indices for *) + (* its arguments that generate the 'elem'-th element of *) + (* the datatype given by 'idx' *) + (* int -> int -> int * int list *) + fun get_cargs idx elem = + let + (* int * interpretation list -> int * int list *) + fun get_cargs_rec (_, []) = + raise REFUTE ("IDT_recursion_interpreter", + "no matching constructor found for datatype element") + | get_cargs_rec (n, x::xs) = + (case get_args x elem of + SOME args => (n, args) + | NONE => get_cargs_rec (n+1, xs)) + in + get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx)) + end + (* computes one entry in 'REC_OPERATORS', and recursively *) + (* all entries needed for it, where 'idx' gives the *) + (* datatype and 'elem' the element of it *) + (* int -> int -> interpretation *) + fun compute_array_entry idx elem = + let + val arr = the (AList.lookup (op =) REC_OPERATORS idx) + val (flag, intr) = Array.sub (arr, elem) + in + if flag then + (* simply return the previously computed result *) + intr + else + (* we have to apply 'intr' to interpretations for all *) + (* recursive arguments *) + let + (* int * int list *) + val (c, args) = get_cargs idx elem + (* find the indices of the constructor's /recursive/ *) + (* arguments *) + val (_, _, constrs) = the (AList.lookup (op =) descr idx) + val (_, dtyps) = List.nth (constrs, c) + val rec_dtyps_args = filter + (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args) + (* map those indices to interpretations *) + val rec_dtyps_intrs = map (fn (dtyp, arg) => + let + val dT = typ_of_dtyp descr typ_assoc dtyp + val consts = make_constants ctxt (typs, []) dT + val arg_i = List.nth (consts, arg) + in + (dtyp, arg_i) + end) rec_dtyps_args + (* takes the dtyp and interpretation of an element, *) + (* and computes the interpretation for the *) + (* corresponding recursive argument *) + fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) = + (* recursive argument is "rec_i params elem" *) + compute_array_entry i (find_index (fn x => x = True) xs) + | rec_intr (Datatype_Aux.DtRec _) (Node _) = + raise REFUTE ("IDT_recursion_interpreter", + "interpretation for IDT is a node") + | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) (Node xs) = + (* recursive argument is something like *) + (* "\<lambda>x::dt1. rec_? params (elem x)" *) + Node (map (rec_intr dt2) xs) + | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) = + raise REFUTE ("IDT_recursion_interpreter", + "interpretation for function dtyp is a leaf") + | rec_intr _ _ = + (* admissibility ensures that every recursive type *) + (* is of the form 'Dt_1 -> ... -> Dt_k -> *) + (* (DtRec i)' *) + raise REFUTE ("IDT_recursion_interpreter", + "non-recursive codomain in recursive dtyp") + (* obtain interpretations for recursive arguments *) + (* interpretation list *) + val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs + (* apply 'intr' to all recursive arguments *) + val result = fold (fn arg_i => fn i => + interpretation_apply (i, arg_i)) arg_intrs intr + (* update 'REC_OPERATORS' *) + val _ = Array.update (arr, elem, (true, result)) + in + result + end + end + val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index)) + (* sanity check: the size of 'IDT' should be 'idt_size' *) + val _ = + if idt_size <> size_of_type ctxt (typs, []) IDT then + raise REFUTE ("IDT_recursion_interpreter", + "unexpected size of IDT (wrong type associated?)") + else () + (* interpretation *) + val rec_op = Node (map_range (compute_array_entry idt_index) idt_size) in - (idx, map (fn (cname, cargs) => - (#1 o interpret thy (typs, []) {maxvars=0, - def_eq=false, next_idx=1, bounds=[], - wellformed=True}) (Const (cname, map (typ_of_dtyp - descr typ_assoc) cargs ---> c_return_typ))) cs) - end) descr - (* associate constructors with corresponding parameters *) - (* (int * (interpretation * interpretation) list) list *) - val (mc_p_intrs, p_intrs') = fold_map - (fn (idx, c_intrs) => fn p_intrs' => - let - val len = length c_intrs - in - ((idx, c_intrs ~~ List.take (p_intrs', len)), - List.drop (p_intrs', len)) - end) mc_intrs p_intrs - (* sanity check: no 'p_intr' may be left afterwards *) - val _ = if p_intrs' <> [] then - raise REFUTE ("IDT_recursion_interpreter", - "more parameter than constructor interpretations") - else () - (* The recursion operator, applied to 'mconstrs_count' *) - (* arguments, is a function that maps every element of the *) - (* inductive datatype to an element of some result type. *) - (* Recursion operators for mutually recursive IDTs are *) - (* translated simultaneously. *) - (* Since the order on datatype elements is given by an *) - (* order on constructors (and then by the order on *) - (* argument tuples), we can simply copy corresponding *) - (* subtrees from 'p_intrs', in the order in which they are *) - (* given. *) - (* interpretation * interpretation -> interpretation list *) - fun ci_pi (Leaf xs, pi) = - (* if the constructor does not match the arguments to a *) - (* defined element of the IDT, the corresponding value *) - (* of the parameter must be ignored *) - if List.exists (equal True) xs then [pi] else [] - | ci_pi (Node xs, Node ys) = - maps ci_pi (xs ~~ ys) - | ci_pi (Node _, Leaf _) = - raise REFUTE ("IDT_recursion_interpreter", - "constructor takes more arguments than the " ^ - "associated parameter") - (* (int * interpretation list) list *) - val rec_operators = map (fn (idx, c_p_intrs) => - (idx, maps ci_pi c_p_intrs)) mc_p_intrs - (* sanity check: every recursion operator must provide as *) - (* many values as the corresponding datatype *) - (* has elements *) - val _ = map (fn (idx, intrs) => - let - val T = typ_of_dtyp descr typ_assoc - (Datatype_Aux.DtRec idx) - in - if length intrs <> size_of_type thy (typs, []) T then - raise REFUTE ("IDT_recursion_interpreter", - "wrong number of interpretations for rec. operator") - else () - end) rec_operators - (* For non-recursive datatypes, we are pretty much done at *) - (* this point. For recursive datatypes however, we still *) - (* need to apply the interpretations in 'rec_operators' to *) - (* (recursively obtained) interpretations for recursive *) - (* constructor arguments. To do so more efficiently, we *) - (* copy 'rec_operators' into arrays first. Each Boolean *) - (* indicates whether the recursive arguments have been *) - (* considered already. *) - (* (int * (bool * interpretation) Array.array) list *) - val REC_OPERATORS = map (fn (idx, intrs) => - (idx, Array.fromList (map (pair false) intrs))) - rec_operators - (* takes an interpretation, and if some leaf of this *) - (* interpretation is the 'elem'-th element of the type, *) - (* the indices of the arguments leading to this leaf are *) - (* returned *) - (* interpretation -> int -> int list option *) - fun get_args (Leaf xs) elem = - if find_index (fn x => x = True) xs = elem then - SOME [] - else - NONE - | get_args (Node xs) elem = - let - (* interpretation list * int -> int list option *) - fun search ([], _) = - NONE - | search (x::xs, n) = - (case get_args x elem of - SOME result => SOME (n::result) - | NONE => search (xs, n+1)) - in - search (xs, 0) + SOME (rec_op, model', args') end - (* returns the index of the constructor and indices for *) - (* its arguments that generate the 'elem'-th element of *) - (* the datatype given by 'idx' *) - (* int -> int -> int * int list *) - fun get_cargs idx elem = - let - (* int * interpretation list -> int * int list *) - fun get_cargs_rec (_, []) = - raise REFUTE ("IDT_recursion_interpreter", - "no matching constructor found for datatype element") - | get_cargs_rec (n, x::xs) = - (case get_args x elem of - SOME args => (n, args) - | NONE => get_cargs_rec (n+1, xs)) - in - get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx)) - end - (* computes one entry in 'REC_OPERATORS', and recursively *) - (* all entries needed for it, where 'idx' gives the *) - (* datatype and 'elem' the element of it *) - (* int -> int -> interpretation *) - fun compute_array_entry idx elem = - let - val arr = the (AList.lookup (op =) REC_OPERATORS idx) - val (flag, intr) = Array.sub (arr, elem) - in - if flag then - (* simply return the previously computed result *) - intr - else - (* we have to apply 'intr' to interpretations for all *) - (* recursive arguments *) - let - (* int * int list *) - val (c, args) = get_cargs idx elem - (* find the indices of the constructor's /recursive/ *) - (* arguments *) - val (_, _, constrs) = the (AList.lookup (op =) descr idx) - val (_, dtyps) = List.nth (constrs, c) - val rec_dtyps_args = filter - (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args) - (* map those indices to interpretations *) - val rec_dtyps_intrs = map (fn (dtyp, arg) => - let - val dT = typ_of_dtyp descr typ_assoc dtyp - val consts = make_constants thy (typs, []) dT - val arg_i = List.nth (consts, arg) - in - (dtyp, arg_i) - end) rec_dtyps_args - (* takes the dtyp and interpretation of an element, *) - (* and computes the interpretation for the *) - (* corresponding recursive argument *) - fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) = - (* recursive argument is "rec_i params elem" *) - compute_array_entry i (find_index (fn x => x = True) xs) - | rec_intr (Datatype_Aux.DtRec _) (Node _) = - raise REFUTE ("IDT_recursion_interpreter", - "interpretation for IDT is a node") - | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) - (Node xs) = - (* recursive argument is something like *) - (* "\<lambda>x::dt1. rec_? params (elem x)" *) - Node (map (rec_intr dt2) xs) - | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) - (Leaf _) = - raise REFUTE ("IDT_recursion_interpreter", - "interpretation for function dtyp is a leaf") - | rec_intr _ _ = - (* admissibility ensures that every recursive type *) - (* is of the form 'Dt_1 -> ... -> Dt_k -> *) - (* (DtRec i)' *) - raise REFUTE ("IDT_recursion_interpreter", - "non-recursive codomain in recursive dtyp") - (* obtain interpretations for recursive arguments *) - (* interpretation list *) - val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs - (* apply 'intr' to all recursive arguments *) - val result = fold (fn arg_i => fn i => - interpretation_apply (i, arg_i)) arg_intrs intr - (* update 'REC_OPERATORS' *) - val _ = Array.update (arr, elem, (true, result)) - in - result - end - end - val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index)) - (* sanity check: the size of 'IDT' should be 'idt_size' *) - val _ = if idt_size <> size_of_type thy (typs, []) IDT then - raise REFUTE ("IDT_recursion_interpreter", - "unexpected size of IDT (wrong type associated?)") - else () - (* interpretation *) - val rec_op = Node (map_range (compute_array_entry idt_index) idt_size) - in - SOME (rec_op, model', args') end - end - else - NONE (* not a recursion operator of this datatype *) - ) (Datatype.get_all thy) NONE + else + NONE (* not a recursion operator of this datatype *) + ) (Datatype.get_all thy) NONE | _ => (* head of term is not a constant *) - NONE; + NONE + end; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) - - fun set_interpreter thy model args t = +fun set_interpreter ctxt model args t = let val (typs, terms) = model in case AList.lookup (op =) terms t of SOME intr => - (* return an existing interpretation *) - SOME (intr, model, args) + (* return an existing interpretation *) + SOME (intr, model, args) | NONE => - (case t of - (* 'Collect' == identity *) - Const (@{const_name Collect}, _) $ t1 => - SOME (interpret thy model args t1) - | Const (@{const_name Collect}, _) => - SOME (interpret thy model args (eta_expand t 1)) - (* 'op :' == application *) - | Const (@{const_name Set.member}, _) $ t1 $ t2 => - SOME (interpret thy model args (t2 $ t1)) - | Const (@{const_name Set.member}, _) $ t1 => - SOME (interpret thy model args (eta_expand t 1)) - | Const (@{const_name Set.member}, _) => - SOME (interpret thy model args (eta_expand t 2)) - | _ => NONE) + (case t of + (* 'Collect' == identity *) + Const (@{const_name Collect}, _) $ t1 => + SOME (interpret ctxt model args t1) + | Const (@{const_name Collect}, _) => + SOME (interpret ctxt model args (eta_expand t 1)) + (* 'op :' == application *) + | Const (@{const_name Set.member}, _) $ t1 $ t2 => + SOME (interpret ctxt model args (t2 $ t1)) + | Const (@{const_name Set.member}, _) $ t1 => + SOME (interpret ctxt model args (eta_expand t 1)) + | Const (@{const_name Set.member}, _) => + SOME (interpret ctxt model args (eta_expand t 2)) + | _ => NONE) end; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* only an optimization: 'card' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - (* only an optimization: 'card' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) - - fun Finite_Set_card_interpreter thy model args t = - case t of - Const (@{const_name Finite_Set.card}, - Type ("fun", [Type ("fun", [T, @{typ bool}]), - @{typ nat}])) => +fun Finite_Set_card_interpreter ctxt model args t = + case t of + Const (@{const_name Finite_Set.card}, + Type ("fun", [Type ("fun", [T, @{typ bool}]), @{typ nat}])) => let (* interpretation -> int *) fun number_of_elements (Node xs) = @@ -2668,9 +2575,9 @@ "interpretation for set type does not yield a Boolean")) xs 0 | number_of_elements (Leaf _) = - raise REFUTE ("Finite_Set_card_interpreter", - "interpretation for set type is a leaf") - val size_of_nat = size_of_type thy model (@{typ nat}) + raise REFUTE ("Finite_Set_card_interpreter", + "interpretation for set type is a leaf") + val size_of_nat = size_of_type ctxt model (@{typ nat}) (* takes an interpretation for a set and returns an interpretation *) (* for a 'nat' denoting the set's cardinality *) (* interpretation -> interpretation *) @@ -2678,62 +2585,54 @@ let val n = number_of_elements i in - if n<size_of_nat then + if n < size_of_nat then Leaf ((replicate n False) @ True :: (replicate (size_of_nat-n-1) False)) else Leaf (replicate size_of_nat False) end val set_constants = - make_constants thy model (Type ("fun", [T, HOLogic.boolT])) + make_constants ctxt model (Type ("fun", [T, HOLogic.boolT])) in SOME (Node (map card set_constants), model, args) end - | _ => - NONE; + | _ => NONE; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* only an optimization: 'finite' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - (* only an optimization: 'finite' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) - - fun Finite_Set_finite_interpreter thy model args t = - case t of - Const (@{const_name Finite_Set.finite}, - Type ("fun", [Type ("fun", [T, @{typ bool}]), - @{typ bool}])) $ _ => +fun Finite_Set_finite_interpreter ctxt model args t = + case t of + Const (@{const_name Finite_Set.finite}, + Type ("fun", [Type ("fun", [T, @{typ bool}]), + @{typ bool}])) $ _ => (* we only consider finite models anyway, hence EVERY set is *) (* "finite" *) SOME (TT, model, args) - | Const (@{const_name Finite_Set.finite}, - Type ("fun", [Type ("fun", [T, @{typ bool}]), - @{typ bool}])) => + | Const (@{const_name Finite_Set.finite}, + Type ("fun", [Type ("fun", [T, @{typ bool}]), + @{typ bool}])) => let val size_of_set = - size_of_type thy model (Type ("fun", [T, HOLogic.boolT])) + size_of_type ctxt model (Type ("fun", [T, HOLogic.boolT])) in (* we only consider finite models anyway, hence EVERY set is *) (* "finite" *) SOME (Node (replicate size_of_set TT), model, args) end - | _ => - NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* only an optimization: 'less' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'less' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun Nat_less_interpreter thy model args t = - case t of - Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat}, +fun Nat_less_interpreter ctxt model args t = + case t of + Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat}, Type ("fun", [@{typ nat}, @{typ bool}])])) => let - val size_of_nat = size_of_type thy model (@{typ nat}) + val size_of_nat = size_of_type ctxt model (@{typ nat}) (* the 'n'-th nat is not less than the first 'n' nats, while it *) (* is less than the remaining 'size_of_nat - n' nats *) (* int -> interpretation *) @@ -2741,22 +2640,18 @@ in SOME (Node (map less (1 upto size_of_nat)), model, args) end - | _ => - NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* only an optimization: 'plus' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'plus' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun Nat_plus_interpreter thy model args t = - case t of - Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat}, +fun Nat_plus_interpreter ctxt model args t = + case t of + Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat}, Type ("fun", [@{typ nat}, @{typ nat}])])) => let - val size_of_nat = size_of_type thy model (@{typ nat}) + val size_of_nat = size_of_type ctxt model (@{typ nat}) (* int -> int -> interpretation *) fun plus m n = let @@ -2772,22 +2667,18 @@ SOME (Node (map_range (fn m => Node (map_range (plus m) size_of_nat)) size_of_nat), model, args) end - | _ => - NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* only an optimization: 'minus' could in principle be interpreted *) - (* with interpreters available already (using its definition), but the *) - (* code below is more efficient *) +(* only an optimization: 'minus' could in principle be interpreted *) +(* with interpreters available already (using its definition), but the *) +(* code below is more efficient *) - fun Nat_minus_interpreter thy model args t = - case t of - Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat}, +fun Nat_minus_interpreter ctxt model args t = + case t of + Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat}, Type ("fun", [@{typ nat}, @{typ nat}])])) => let - val size_of_nat = size_of_type thy model (@{typ nat}) + val size_of_nat = size_of_type ctxt model (@{typ nat}) (* int -> int -> interpretation *) fun minus m n = let @@ -2800,22 +2691,18 @@ SOME (Node (map_range (fn m => Node (map_range (minus m) size_of_nat)) size_of_nat), model, args) end - | _ => - NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* only an optimization: 'times' could in principle be interpreted *) - (* with interpreters available already (using its definition), but the *) - (* code below is more efficient *) +(* only an optimization: 'times' could in principle be interpreted *) +(* with interpreters available already (using its definition), but the *) +(* code below is more efficient *) - fun Nat_times_interpreter thy model args t = - case t of - Const (@{const_name Groups.times}, Type ("fun", [@{typ nat}, +fun Nat_times_interpreter ctxt model args t = + case t of + Const (@{const_name Groups.times}, Type ("fun", [@{typ nat}, Type ("fun", [@{typ nat}, @{typ nat}])])) => let - val size_of_nat = size_of_type thy model (@{typ nat}) + val size_of_nat = size_of_type ctxt model (@{typ nat}) (* nat -> nat -> interpretation *) fun mult m n = let @@ -2831,25 +2718,22 @@ SOME (Node (map_range (fn m => Node (map_range (mult m) size_of_nat)) size_of_nat), model, args) end - | _ => - NONE; - - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) + | _ => NONE; - (* only an optimization: 'append' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'append' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun List_append_interpreter thy model args t = - case t of - Const (@{const_name List.append}, Type ("fun", [Type ("List.list", [T]), Type ("fun", +fun List_append_interpreter ctxt model args t = + case t of + Const (@{const_name List.append}, Type ("fun", [Type ("List.list", [T]), Type ("fun", [Type ("List.list", [_]), Type ("List.list", [_])])])) => let - val size_elem = size_of_type thy model T - val size_list = size_of_type thy model (Type ("List.list", [T])) + val size_elem = size_of_type ctxt model T + val size_list = size_of_type ctxt model (Type ("List.list", [T])) (* maximal length of lists; 0 if we only consider the empty list *) - val list_length = let + val list_length = + let (* int -> int -> int -> int *) fun list_length_acc len lists total = if lists = total then @@ -2893,12 +2777,12 @@ in case offsets' of [] => - (* we're at the last node in the tree; the next value *) - (* won't be used anyway *) - (assoc, ([], 0)) + (* we're at the last node in the tree; the next value *) + (* won't be used anyway *) + (assoc, ([], 0)) | off'::offs' => - (* go to next sibling node *) - (assoc, (offs', off' + 1)) + (* go to next sibling node *) + (assoc, (offs', off' + 1)) end end) elements ([], 0) (* we also need the reverse association (from length/offset to *) @@ -2911,341 +2795,328 @@ val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m) val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n) val len_elem = len_m + len_n - val off_elem = off_m * power (size_elem, len_n) + off_n + val off_elem = off_m * Integer.pow len_n size_elem + off_n in - case AList.lookup op= lenoff'_lists (len_elem, off_elem) of + case AList.lookup op= lenoff'_lists (len_elem, off_elem) of NONE => - (* undefined *) - Leaf (replicate size_list False) + (* undefined *) + Leaf (replicate size_list False) | SOME element => - Leaf ((replicate element False) @ True :: - (replicate (size_list - element - 1) False)) + Leaf ((replicate element False) @ True :: + (replicate (size_list - element - 1) False)) end in SOME (Node (map (fn m => Node (map (append m) elements)) elements), model, args) end - | _ => - NONE; + | _ => NONE; (* UNSOUND - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) - - (* only an optimization: 'lfp' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'lfp' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun lfp_interpreter thy model args t = - case t of - Const (@{const_name lfp}, Type ("fun", [Type ("fun", - [Type ("fun", [T, @{typ bool}]), - Type ("fun", [_, @{typ bool}])]), - Type ("fun", [_, @{typ bool}])])) => +fun lfp_interpreter ctxt model args t = + case t of + Const (@{const_name lfp}, Type ("fun", [Type ("fun", + [Type ("fun", [T, @{typ bool}]), + Type ("fun", [_, @{typ bool}])]), + Type ("fun", [_, @{typ bool}])])) => let - val size_elem = size_of_type thy model T + val size_elem = size_of_type ctxt model T (* the universe (i.e. the set that contains every element) *) val i_univ = Node (replicate size_elem TT) (* all sets with elements from type 'T' *) val i_sets = - make_constants thy model (Type ("fun", [T, HOLogic.boolT])) + make_constants ctxt model (Type ("fun", [T, HOLogic.boolT])) (* all functions that map sets to sets *) - val i_funs = make_constants thy model (Type ("fun", + val i_funs = make_constants ctxt model (Type ("fun", [Type ("fun", [T, @{typ bool}]), Type ("fun", [T, @{typ bool}])])) (* "lfp(f) == Inter({u. f(u) <= u})" *) (* interpretation * interpretation -> bool *) fun is_subset (Node subs, Node sups) = - forall (fn (sub, sup) => (sub = FF) orelse (sup = TT)) - (subs ~~ sups) + forall (fn (sub, sup) => (sub = FF) orelse (sup = TT)) (subs ~~ sups) | is_subset (_, _) = - raise REFUTE ("lfp_interpreter", - "is_subset: interpretation for set is not a node") + raise REFUTE ("lfp_interpreter", + "is_subset: interpretation for set is not a node") (* interpretation * interpretation -> interpretation *) fun intersection (Node xs, Node ys) = - Node (map (fn (x, y) => if x=TT andalso y=TT then TT else FF) - (xs ~~ ys)) + Node (map (fn (x, y) => if x=TT andalso y=TT then TT else FF) + (xs ~~ ys)) | intersection (_, _) = - raise REFUTE ("lfp_interpreter", - "intersection: interpretation for set is not a node") + raise REFUTE ("lfp_interpreter", + "intersection: interpretation for set is not a node") (* interpretation -> interpretaion *) fun lfp (Node resultsets) = - fold (fn (set, resultset) => fn acc => - if is_subset (resultset, set) then - intersection (acc, set) - else - acc) (i_sets ~~ resultsets) i_univ + fold (fn (set, resultset) => fn acc => + if is_subset (resultset, set) then + intersection (acc, set) + else + acc) (i_sets ~~ resultsets) i_univ | lfp _ = - raise REFUTE ("lfp_interpreter", - "lfp: interpretation for function is not a node") + raise REFUTE ("lfp_interpreter", + "lfp: interpretation for function is not a node") in SOME (Node (map lfp i_funs), model, args) end - | _ => - NONE; + | _ => NONE; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) - - (* only an optimization: 'gfp' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'gfp' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun gfp_interpreter thy model args t = - case t of - Const (@{const_name gfp}, Type ("fun", [Type ("fun", - [Type ("fun", [T, @{typ bool}]), - Type ("fun", [_, @{typ bool}])]), - Type ("fun", [_, @{typ bool}])])) => - let - val size_elem = size_of_type thy model T - (* the universe (i.e. the set that contains every element) *) - val i_univ = Node (replicate size_elem TT) - (* all sets with elements from type 'T' *) - val i_sets = - make_constants thy model (Type ("fun", [T, HOLogic.boolT])) - (* all functions that map sets to sets *) - val i_funs = make_constants thy model (Type ("fun", - [Type ("fun", [T, HOLogic.boolT]), - Type ("fun", [T, HOLogic.boolT])])) - (* "gfp(f) == Union({u. u <= f(u)})" *) - (* interpretation * interpretation -> bool *) - fun is_subset (Node subs, Node sups) = - forall (fn (sub, sup) => (sub = FF) orelse (sup = TT)) - (subs ~~ sups) - | is_subset (_, _) = - raise REFUTE ("gfp_interpreter", - "is_subset: interpretation for set is not a node") - (* interpretation * interpretation -> interpretation *) - fun union (Node xs, Node ys) = +fun gfp_interpreter ctxt model args t = + case t of + Const (@{const_name gfp}, Type ("fun", [Type ("fun", + [Type ("fun", [T, @{typ bool}]), + Type ("fun", [_, @{typ bool}])]), + Type ("fun", [_, @{typ bool}])])) => + let + val size_elem = size_of_type ctxt model T + (* the universe (i.e. the set that contains every element) *) + val i_univ = Node (replicate size_elem TT) + (* all sets with elements from type 'T' *) + val i_sets = + make_constants ctxt model (Type ("fun", [T, HOLogic.boolT])) + (* all functions that map sets to sets *) + val i_funs = make_constants ctxt model (Type ("fun", + [Type ("fun", [T, HOLogic.boolT]), + Type ("fun", [T, HOLogic.boolT])])) + (* "gfp(f) == Union({u. u <= f(u)})" *) + (* interpretation * interpretation -> bool *) + fun is_subset (Node subs, Node sups) = + forall (fn (sub, sup) => (sub = FF) orelse (sup = TT)) + (subs ~~ sups) + | is_subset (_, _) = + raise REFUTE ("gfp_interpreter", + "is_subset: interpretation for set is not a node") + (* interpretation * interpretation -> interpretation *) + fun union (Node xs, Node ys) = Node (map (fn (x,y) => if x=TT orelse y=TT then TT else FF) (xs ~~ ys)) - | union (_, _) = - raise REFUTE ("gfp_interpreter", - "union: interpretation for set is not a node") - (* interpretation -> interpretaion *) - fun gfp (Node resultsets) = - fold (fn (set, resultset) => fn acc => - if is_subset (set, resultset) then - union (acc, set) - else - acc) (i_sets ~~ resultsets) i_univ - | gfp _ = + | union (_, _) = + raise REFUTE ("gfp_interpreter", + "union: interpretation for set is not a node") + (* interpretation -> interpretaion *) + fun gfp (Node resultsets) = + fold (fn (set, resultset) => fn acc => + if is_subset (set, resultset) then + union (acc, set) + else + acc) (i_sets ~~ resultsets) i_univ + | gfp _ = raise REFUTE ("gfp_interpreter", "gfp: interpretation for function is not a node") - in - SOME (Node (map gfp i_funs), model, args) - end - | _ => - NONE; + in + SOME (Node (map gfp i_funs), model, args) + end + | _ => NONE; *) - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) - - (* only an optimization: 'fst' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) +(* only an optimization: 'fst' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - fun Product_Type_fst_interpreter thy model args t = - case t of - Const (@{const_name fst}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) => +fun Product_Type_fst_interpreter ctxt model args t = + case t of + Const (@{const_name fst}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) => let - val constants_T = make_constants thy model T - val size_U = size_of_type thy model U + val constants_T = make_constants ctxt model T + val size_U = size_of_type ctxt model U in SOME (Node (maps (replicate size_U) constants_T), model, args) end - | _ => - NONE; + | _ => NONE; - (* theory -> model -> arguments -> Term.term -> - (interpretation * model * arguments) option *) +(* only an optimization: 'snd' could in principle be interpreted with *) +(* interpreters available already (using its definition), but the code *) +(* below is more efficient *) - (* only an optimization: 'snd' could in principle be interpreted with *) - (* interpreters available already (using its definition), but the code *) - (* below is more efficient *) - - fun Product_Type_snd_interpreter thy model args t = - case t of - Const (@{const_name snd}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) => +fun Product_Type_snd_interpreter ctxt model args t = + case t of + Const (@{const_name snd}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) => let - val size_T = size_of_type thy model T - val constants_U = make_constants thy model U + val size_T = size_of_type ctxt model T + val constants_U = make_constants ctxt model U in SOME (Node (flat (replicate size_T constants_U)), model, args) end - | _ => - NONE; + | _ => NONE; (* ------------------------------------------------------------------------- *) (* PRINTERS *) (* ------------------------------------------------------------------------- *) - (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> - Term.term option *) - - fun stlc_printer thy model T intr assignment = +fun stlc_printer ctxt model T intr assignment = let (* string -> string *) fun strip_leading_quote s = (implode o (fn [] => [] | x::xs => if x="'" then xs else x::xs) o explode) s (* Term.typ -> string *) - fun string_of_typ (Type (s, _)) = s - | string_of_typ (TFree (x, _)) = strip_leading_quote x + fun string_of_typ (Type (s, _)) = s + | string_of_typ (TFree (x, _)) = strip_leading_quote x | string_of_typ (TVar ((x,i), _)) = - strip_leading_quote x ^ string_of_int i + strip_leading_quote x ^ string_of_int i (* interpretation -> int *) fun index_from_interpretation (Leaf xs) = - find_index (PropLogic.eval assignment) xs + find_index (PropLogic.eval assignment) xs | index_from_interpretation _ = - raise REFUTE ("stlc_printer", - "interpretation for ground type is not a leaf") + raise REFUTE ("stlc_printer", + "interpretation for ground type is not a leaf") in case T of Type ("fun", [T1, T2]) => - let - (* create all constants of type 'T1' *) - val constants = make_constants thy model T1 - (* interpretation list *) - val results = (case intr of - Node xs => xs - | _ => raise REFUTE ("stlc_printer", - "interpretation for function type is a leaf")) - (* Term.term list *) - val pairs = map (fn (arg, result) => - HOLogic.mk_prod - (print thy model T1 arg assignment, - print thy model T2 result assignment)) - (constants ~~ results) - (* Term.typ *) - val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) - val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT - (* Term.term *) - val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT) - val HOLogic_insert = - Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) - in - SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set) - end - | Type ("prop", []) => - (case index_from_interpretation intr of - ~1 => SOME (HOLogic.mk_Trueprop (Const (@{const_name undefined}, HOLogic.boolT))) - | 0 => SOME (HOLogic.mk_Trueprop HOLogic.true_const) - | 1 => SOME (HOLogic.mk_Trueprop HOLogic.false_const) - | _ => raise REFUTE ("stlc_interpreter", - "illegal interpretation for a propositional value")) - | Type _ => if index_from_interpretation intr = (~1) then - SOME (Const (@{const_name undefined}, T)) - else - SOME (Const (string_of_typ T ^ - string_of_int (index_from_interpretation intr), T)) - | TFree _ => if index_from_interpretation intr = (~1) then - SOME (Const (@{const_name undefined}, T)) - else - SOME (Const (string_of_typ T ^ - string_of_int (index_from_interpretation intr), T)) - | TVar _ => if index_from_interpretation intr = (~1) then - SOME (Const (@{const_name undefined}, T)) - else - SOME (Const (string_of_typ T ^ - string_of_int (index_from_interpretation intr), T)) + let + (* create all constants of type 'T1' *) + val constants = make_constants ctxt model T1 + (* interpretation list *) + val results = + (case intr of + Node xs => xs + | _ => raise REFUTE ("stlc_printer", + "interpretation for function type is a leaf")) + (* Term.term list *) + val pairs = map (fn (arg, result) => + HOLogic.mk_prod + (print ctxt model T1 arg assignment, + print ctxt model T2 result assignment)) + (constants ~~ results) + (* Term.typ *) + val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) + val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT + (* Term.term *) + val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT) + val HOLogic_insert = + Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) + in + SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set) + end + | Type ("prop", []) => + (case index_from_interpretation intr of + ~1 => SOME (HOLogic.mk_Trueprop (Const (@{const_name undefined}, HOLogic.boolT))) + | 0 => SOME (HOLogic.mk_Trueprop HOLogic.true_const) + | 1 => SOME (HOLogic.mk_Trueprop HOLogic.false_const) + | _ => raise REFUTE ("stlc_interpreter", + "illegal interpretation for a propositional value")) + | Type _ => + if index_from_interpretation intr = (~1) then + SOME (Const (@{const_name undefined}, T)) + else + SOME (Const (string_of_typ T ^ + string_of_int (index_from_interpretation intr), T)) + | TFree _ => + if index_from_interpretation intr = (~1) then + SOME (Const (@{const_name undefined}, T)) + else + SOME (Const (string_of_typ T ^ + string_of_int (index_from_interpretation intr), T)) + | TVar _ => + if index_from_interpretation intr = (~1) then + SOME (Const (@{const_name undefined}, T)) + else + SOME (Const (string_of_typ T ^ + string_of_int (index_from_interpretation intr), T)) end; - (* theory -> model -> Term.typ -> interpretation -> (int -> bool) -> - Term.term option *) - - fun IDT_printer thy model T intr assignment = +fun IDT_printer ctxt model T intr assignment = + let + val thy = ProofContext.theory_of ctxt + in (case T of Type (s, Ts) => - (case Datatype.get_info thy s of - SOME info => (* inductive datatype *) - let - val (typs, _) = model - val index = #index info - val descr = #descr info - val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) - val typ_assoc = dtyps ~~ Ts - (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) - val _ = if Library.exists (fn d => - case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps - then - raise REFUTE ("IDT_printer", "datatype argument (for type " ^ - Syntax.string_of_typ_global thy (Type (s, Ts)) ^ ") is not a variable") - else () - (* the index of the element in the datatype *) - val element = (case intr of - Leaf xs => find_index (PropLogic.eval assignment) xs - | Node _ => raise REFUTE ("IDT_printer", - "interpretation is not a leaf")) - in - if element < 0 then - SOME (Const (@{const_name undefined}, Type (s, Ts))) - else let - (* takes a datatype constructor, and if for some arguments this *) - (* constructor generates the datatype's element that is given by *) - (* 'element', returns the constructor (as a term) as well as the *) - (* indices of the arguments *) - fun get_constr_args (cname, cargs) = - let - val cTerm = Const (cname, - map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts)) - val (iC, _, _) = interpret thy (typs, []) {maxvars=0, - def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm - (* interpretation -> int list option *) - fun get_args (Leaf xs) = - if find_index (fn x => x = True) xs = element then - SOME [] - else - NONE - | get_args (Node xs) = - let - (* interpretation * int -> int list option *) - fun search ([], _) = - NONE - | search (x::xs, n) = - (case get_args x of - SOME result => SOME (n::result) - | NONE => search (xs, n+1)) - in - search (xs, 0) - end - in - Option.map (fn args => (cTerm, cargs, args)) (get_args iC) - end - val (cTerm, cargs, args) = - (* we could speed things up by computing the correct *) - (* constructor directly (rather than testing all *) - (* constructors), based on the order in which constructors *) - (* generate elements of datatypes; the current implementation *) - (* of 'IDT_printer' however is independent of the internals *) - (* of 'IDT_constructor_interpreter' *) - (case get_first get_constr_args constrs of - SOME x => x - | NONE => raise REFUTE ("IDT_printer", - "no matching constructor found for element " ^ - string_of_int element)) - val argsTerms = map (fn (d, n) => - let - val dT = typ_of_dtyp descr typ_assoc d - (* we only need the n-th element of this list, so there *) - (* might be a more efficient implementation that does not *) - (* generate all constants *) - val consts = make_constants thy (typs, []) dT - in - print thy (typs, []) dT (List.nth (consts, n)) assignment - end) (cargs ~~ args) - in - SOME (list_comb (cTerm, argsTerms)) - end - end - | NONE => (* not an inductive datatype *) + (case Datatype.get_info thy s of + SOME info => (* inductive datatype *) + let + val (typs, _) = model + val index = #index info + val descr = #descr info + val (_, dtyps, constrs) = the (AList.lookup (op =) descr index) + val typ_assoc = dtyps ~~ Ts + (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) + val _ = + if Library.exists (fn d => + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps + then + raise REFUTE ("IDT_printer", "datatype argument (for type " ^ + Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable") + else () + (* the index of the element in the datatype *) + val element = + (case intr of + Leaf xs => find_index (PropLogic.eval assignment) xs + | Node _ => raise REFUTE ("IDT_printer", + "interpretation is not a leaf")) + in + if element < 0 then + SOME (Const (@{const_name undefined}, Type (s, Ts))) + else + let + (* takes a datatype constructor, and if for some arguments this *) + (* constructor generates the datatype's element that is given by *) + (* 'element', returns the constructor (as a term) as well as the *) + (* indices of the arguments *) + fun get_constr_args (cname, cargs) = + let + val cTerm = Const (cname, + map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts)) + val (iC, _, _) = interpret ctxt (typs, []) {maxvars=0, + def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm + (* interpretation -> int list option *) + fun get_args (Leaf xs) = + if find_index (fn x => x = True) xs = element then + SOME [] + else + NONE + | get_args (Node xs) = + let + (* interpretation * int -> int list option *) + fun search ([], _) = + NONE + | search (x::xs, n) = + (case get_args x of + SOME result => SOME (n::result) + | NONE => search (xs, n+1)) + in + search (xs, 0) + end + in + Option.map (fn args => (cTerm, cargs, args)) (get_args iC) + end + val (cTerm, cargs, args) = + (* we could speed things up by computing the correct *) + (* constructor directly (rather than testing all *) + (* constructors), based on the order in which constructors *) + (* generate elements of datatypes; the current implementation *) + (* of 'IDT_printer' however is independent of the internals *) + (* of 'IDT_constructor_interpreter' *) + (case get_first get_constr_args constrs of + SOME x => x + | NONE => raise REFUTE ("IDT_printer", + "no matching constructor found for element " ^ + string_of_int element)) + val argsTerms = map (fn (d, n) => + let + val dT = typ_of_dtyp descr typ_assoc d + (* we only need the n-th element of this list, so there *) + (* might be a more efficient implementation that does not *) + (* generate all constants *) + val consts = make_constants ctxt (typs, []) dT + in + print ctxt (typs, []) dT (List.nth (consts, n)) assignment + end) (cargs ~~ args) + in + SOME (list_comb (cTerm, argsTerms)) + end + end + | NONE => (* not an inductive datatype *) + NONE) + | _ => (* a (free or schematic) type variable *) NONE) - | _ => (* a (free or schematic) type variable *) - NONE); + end; (* ------------------------------------------------------------------------- *) @@ -3260,28 +3131,71 @@ (* subterms that are then passed to other interpreters! *) (* ------------------------------------------------------------------------- *) - val setup = - add_interpreter "stlc" stlc_interpreter #> - add_interpreter "Pure" Pure_interpreter #> - add_interpreter "HOLogic" HOLogic_interpreter #> - add_interpreter "set" set_interpreter #> - add_interpreter "IDT" IDT_interpreter #> - add_interpreter "IDT_constructor" IDT_constructor_interpreter #> - add_interpreter "IDT_recursion" IDT_recursion_interpreter #> - add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #> - add_interpreter "Finite_Set.finite" Finite_Set_finite_interpreter #> - add_interpreter "Nat_Orderings.less" Nat_less_interpreter #> - add_interpreter "Nat_HOL.plus" Nat_plus_interpreter #> - add_interpreter "Nat_HOL.minus" Nat_minus_interpreter #> - add_interpreter "Nat_HOL.times" Nat_times_interpreter #> - add_interpreter "List.append" List_append_interpreter #> +val setup = + add_interpreter "stlc" stlc_interpreter #> + add_interpreter "Pure" Pure_interpreter #> + add_interpreter "HOLogic" HOLogic_interpreter #> + add_interpreter "set" set_interpreter #> + add_interpreter "IDT" IDT_interpreter #> + add_interpreter "IDT_constructor" IDT_constructor_interpreter #> + add_interpreter "IDT_recursion" IDT_recursion_interpreter #> + add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #> + add_interpreter "Finite_Set.finite" Finite_Set_finite_interpreter #> + add_interpreter "Nat_Orderings.less" Nat_less_interpreter #> + add_interpreter "Nat_HOL.plus" Nat_plus_interpreter #> + add_interpreter "Nat_HOL.minus" Nat_minus_interpreter #> + add_interpreter "Nat_HOL.times" Nat_times_interpreter #> + add_interpreter "List.append" List_append_interpreter #> (* UNSOUND - add_interpreter "lfp" lfp_interpreter #> - add_interpreter "gfp" gfp_interpreter #> + add_interpreter "lfp" lfp_interpreter #> + add_interpreter "gfp" gfp_interpreter #> *) - add_interpreter "Product_Type.fst" Product_Type_fst_interpreter #> - add_interpreter "Product_Type.snd" Product_Type_snd_interpreter #> - add_printer "stlc" stlc_printer #> - add_printer "IDT" IDT_printer; + add_interpreter "Product_Type.fst" Product_Type_fst_interpreter #> + add_interpreter "Product_Type.snd" Product_Type_snd_interpreter #> + add_printer "stlc" stlc_printer #> + add_printer "IDT" IDT_printer; + + + +(** outer syntax commands 'refute' and 'refute_params' **) + +(* argument parsing *) + +(*optional list of arguments of the form [name1=value1, name2=value2, ...]*) + +val scan_parm = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- Parse.name) "true") +val scan_parms = Scan.optional (Parse.$$$ "[" |-- Parse.list scan_parm --| Parse.$$$ "]") []; + + +(* 'refute' command *) -end (* structure Refute *) +val _ = + Outer_Syntax.improper_command "refute" + "try to find a model that refutes a given subgoal" Keyword.diag + (scan_parms -- Scan.optional Parse.nat 1 >> + (fn (parms, i) => + Toplevel.keep (fn state => + let + val ctxt = Toplevel.context_of state; + val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state); + in refute_goal ctxt parms st i end))); + + +(* 'refute_params' command *) + +val _ = + Outer_Syntax.command "refute_params" + "show/store default parameters for the 'refute' command" Keyword.thy_decl + (scan_parms >> (fn parms => + Toplevel.theory (fn thy => + let + val thy' = fold set_default_param parms thy; + val output = + (case get_default_params (ProofContext.init_global thy') of + [] => "none" + | new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults)); + val _ = writeln ("Default parameters for 'refute':\n" ^ output); + in thy' end))); + +end; +
--- a/src/HOL/Tools/refute_isar.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/HOL/Tools/refute_isar.ML Fri Sep 03 13:54:39 2010 +0200 @@ -2,49 +2,6 @@ Author: Tjark Weber Copyright 2003-2007 -Outer syntax commands 'refute' and 'refute_params'. -*) - -structure Refute_Isar: sig end = -struct - -(* argument parsing *) - -(*optional list of arguments of the form [name1=value1, name2=value2, ...]*) - -val scan_parm = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- Parse.name) "true") -val scan_parms = Scan.optional (Parse.$$$ "[" |-- Parse.list scan_parm --| Parse.$$$ "]") []; - - -(* 'refute' command *) - -val _ = - Outer_Syntax.improper_command "refute" - "try to find a model that refutes a given subgoal" Keyword.diag - (scan_parms -- Scan.optional Parse.nat 1 >> - (fn (parms, i) => - Toplevel.keep (fn state => - let - val ctxt = Toplevel.context_of state; - val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state); - in Refute.refute_goal ctxt parms st i end))); - - -(* 'refute_params' command *) - -val _ = - Outer_Syntax.command "refute_params" - "show/store default parameters for the 'refute' command" Keyword.thy_decl - (scan_parms >> (fn parms => - Toplevel.theory (fn thy => - let - val thy' = fold Refute.set_default_param parms thy; - val output = - (case Refute.get_default_params thy' of - [] => "none" - | new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults)); - val _ = writeln ("Default parameters for 'refute':\n" ^ output); - in thy' end))); end;
--- a/src/Pure/General/position.scala Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/General/position.scala Fri Sep 03 13:54:39 2010 +0200 @@ -24,7 +24,7 @@ def unapply(pos: T): Option[Text.Range] = (Offset.unapply(pos), End_Offset.unapply(pos)) match { case (Some(start), Some(stop)) if start <= stop => Some(Text.Range(start, stop)) - case (Some(start), None) => Some(Text.Range(start)) + case (Some(start), None) => Some(Text.Range(start, start + 1)) case _ => None } }
--- a/src/Pure/Isar/attrib.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/Isar/attrib.ML Fri Sep 03 13:54:39 2010 +0200 @@ -392,7 +392,8 @@ (* theory setup *) val _ = Context.>> (Context.map_theory - (register_config show_question_marks_value #> + (register_config Syntax.show_question_marks_value #> + register_config Goal_Display.show_consts_value #> register_config Unify.trace_bound_value #> register_config Unify.search_bound_value #> register_config Unify.trace_simp_value #>
--- a/src/Pure/Isar/proof.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/Isar/proof.ML Fri Sep 03 13:54:39 2010 +0200 @@ -331,7 +331,7 @@ fun pretty_state nr state = let - val {context, facts, mode, goal = _} = current state; + val {context = ctxt, facts, mode, goal = _} = current state; fun levels_up 0 = "" | levels_up 1 = "1 level up" @@ -345,7 +345,7 @@ (case filter_out (fn s => s = "") strs of [] => "" | strs' => enclose " (" ")" (commas strs')); - fun prt_goal (SOME (ctxt, (i, + fun prt_goal (SOME (_, (i, {statement = ((_, pos), _, _), messages, using, goal, before_qed = _, after_qed = _}))) = pretty_facts "using " ctxt (if mode <> Backward orelse null using then NONE else SOME using) @ @@ -357,8 +357,8 @@ | prt_goal NONE = []; val prt_ctxt = - if ! verbose orelse mode = Forward then ProofContext.pretty_context context - else if mode = Backward then ProofContext.pretty_ctxt context + if ! verbose orelse mode = Forward then ProofContext.pretty_context ctxt + else if mode = Backward then ProofContext.pretty_ctxt ctxt else []; in [Pretty.str ("proof (" ^ mode_name mode ^ "): step " ^ string_of_int nr ^ @@ -366,8 +366,8 @@ Pretty.str ""] @ (if null prt_ctxt then [] else prt_ctxt @ [Pretty.str ""]) @ (if ! verbose orelse mode = Forward then - pretty_facts "" context facts @ prt_goal (try find_goal state) - else if mode = Chain then pretty_facts "picking " context facts + pretty_facts "" ctxt facts @ prt_goal (try find_goal state) + else if mode = Chain then pretty_facts "picking " ctxt facts else prt_goal (try find_goal state)) end;
--- a/src/Pure/Isar/proof_context.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/Isar/proof_context.ML Fri Sep 03 13:54:39 2010 +0200 @@ -741,14 +741,14 @@ let val (syms, pos) = Syntax.parse_token ctxt Markup.sort text; val S = Syntax.standard_parse_sort ctxt (syn_of ctxt) (syms, pos) - handle ERROR msg => cat_error msg ("Failed to parse sort" ^ Position.str_of pos) + handle ERROR msg => cat_error msg "Failed to parse sort"; in Type.minimize_sort (tsig_of ctxt) S end; fun parse_typ ctxt text = let val (syms, pos) = Syntax.parse_token ctxt Markup.typ text; val T = Syntax.standard_parse_typ ctxt (syn_of ctxt) (get_sort ctxt) (syms, pos) - handle ERROR msg => cat_error msg ("Failed to parse type" ^ Position.str_of pos); + handle ERROR msg => cat_error msg "Failed to parse type"; in T end; fun parse_term T ctxt text = @@ -756,7 +756,8 @@ val {get_sort, map_const, map_free} = term_context ctxt; val (T', _) = Type_Infer.paramify_dummies T 0; - val (markup, kind) = if T' = propT then (Markup.prop, "proposition") else (Markup.term, "term"); + val (markup, kind) = + if T' = propT then (Markup.prop, "proposition") else (Markup.term, "term"); val (syms, pos) = Syntax.parse_token ctxt markup text; fun check t = (Syntax.check_term ctxt (Type_Infer.constrain T' t); NONE) @@ -764,7 +765,7 @@ val t = Syntax.standard_parse_term (Syntax.pp ctxt) check get_sort map_const map_free ctxt (Type.is_logtype (tsig_of ctxt)) (syn_of ctxt) T' (syms, pos) - handle ERROR msg => cat_error msg ("Failed to parse " ^ kind ^ Position.str_of pos); + handle ERROR msg => cat_error msg ("Failed to parse " ^ kind); in t end;
--- a/src/Pure/PIDE/isar_document.scala Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/PIDE/isar_document.scala Fri Sep 03 13:54:39 2010 +0200 @@ -58,17 +58,23 @@ /* reported positions */ + private val include_pos = Set(Markup.BINDING, Markup.ENTITY, Markup.REPORT, Markup.POSITION) + private val exclude_pos = Set(Markup.LOCATION) + def reported_positions(command_id: Document.Command_ID, message: XML.Elem): Set[Text.Range] = { def reported(set: Set[Text.Range], tree: XML.Tree): Set[Text.Range] = tree match { case XML.Elem(Markup(name, Position.Id_Range(id, range)), body) - if (name == Markup.BINDING || name == Markup.ENTITY || name == Markup.REPORT) && - id == command_id => body.foldLeft(set + range)(reported) - case XML.Elem(_, body) => body.foldLeft(set)(reported) - case XML.Text(_) => set + if include_pos(name) && id == command_id => + body.foldLeft(set + range)(reported) + case XML.Elem(Markup(name, _), body) if !exclude_pos(name) => + body.foldLeft(set)(reported) + case _ => set } - reported(Set.empty, message) ++ Position.Range.unapply(message.markup.properties) + val set = reported(Set.empty, message) + if (set.isEmpty) set ++ Position.Range.unapply(message.markup.properties) + else set } }
--- a/src/Pure/ProofGeneral/preferences.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/ProofGeneral/preferences.ML Fri Sep 03 13:54:39 2010 +0200 @@ -117,7 +117,7 @@ bool_pref show_sorts "show-sorts" "Include sorts in display of Isabelle terms", - bool_pref show_consts + bool_pref show_consts_default "show-consts" "Show types of consts in Isabelle goal display", bool_pref long_names
--- a/src/Pure/display.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/display.ML Fri Sep 03 13:54:39 2010 +0200 @@ -8,7 +8,8 @@ signature BASIC_DISPLAY = sig val goals_limit: int Unsynchronized.ref - val show_consts: bool Unsynchronized.ref + val show_consts_default: bool Unsynchronized.ref + val show_consts: bool Config.T val show_hyps: bool Unsynchronized.ref val show_tags: bool Unsynchronized.ref end; @@ -37,6 +38,7 @@ (** options **) val goals_limit = Goal_Display.goals_limit; +val show_consts_default = Goal_Display.show_consts_default; val show_consts = Goal_Display.show_consts; val show_hyps = Unsynchronized.ref false; (*false: print meta-hypotheses as dots*)
--- a/src/Pure/goal_display.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Pure/goal_display.ML Fri Sep 03 13:54:39 2010 +0200 @@ -8,7 +8,9 @@ signature GOAL_DISPLAY = sig val goals_limit: int Unsynchronized.ref - val show_consts: bool Unsynchronized.ref + val show_consts_default: bool Unsynchronized.ref + val show_consts_value: Config.value Config.T + val show_consts: bool Config.T val pretty_flexpair: Proof.context -> term * term -> Pretty.T val pretty_goals: Proof.context -> {total: bool, main: bool, maxgoals: int} -> thm -> Pretty.T list @@ -19,7 +21,12 @@ struct val goals_limit = Unsynchronized.ref 10; (*max number of goals to print*) -val show_consts = Unsynchronized.ref false; (*true: show consts with types in proof state output*) + +(*true: show consts with types in proof state output*) +val show_consts_default = Unsynchronized.ref false; +val show_consts_value = + Config.declare false "show_consts" (fn _ => Config.Bool (! show_consts_default)); +val show_consts = Config.bool show_consts_value; fun pretty_flexpair ctxt (t, u) = Pretty.block [Syntax.pretty_term ctxt t, Pretty.str " =?=", Pretty.brk 1, Syntax.pretty_term ctxt u]; @@ -104,7 +111,7 @@ else []) else pretty_subgoals As) @ pretty_ffpairs tpairs @ - (if ! show_consts then pretty_consts prop else []) @ + (if Config.get ctxt show_consts then pretty_consts prop else []) @ (if types then pretty_vars prop else []) @ (if sorts then pretty_varsT prop else []); in
--- a/src/Tools/Code/code_haskell.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_haskell.ML Fri Sep 03 13:54:39 2010 +0200 @@ -261,13 +261,31 @@ end; in print_stmt end; +fun mk_name_module reserved module_prefix module_alias program = + let + fun mk_alias name = case module_alias name + of SOME name' => name' + | NONE => name + |> Long_Name.explode + |> map (fn name => (the_single o fst) (Name.variants [name] reserved)) + |> Long_Name.implode; + fun mk_prefix name = case module_prefix + of SOME module_prefix => Long_Name.append module_prefix name + | NONE => name; + val tab = + Symtab.empty + |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name)) + o fst o Code_Namespace.dest_name o fst) + program + in the o Symtab.lookup tab end; + fun haskell_program_of_program labelled_name module_prefix reserved module_alias program = let val reserved = Name.make_context reserved; val mk_name_module = mk_name_module reserved module_prefix module_alias program; fun add_stmt (name, (stmt, deps)) = let - val (module_name, base) = dest_name name; + val (module_name, base) = Code_Namespace.dest_name name; val module_name' = mk_name_module module_name; val mk_name_stmt = yield_singleton Name.variants; fun add_fun upper (nsp_fun, nsp_typ) = @@ -309,14 +327,14 @@ (Graph.get_node program name, Graph.imm_succs program name)) (Graph.strong_conn program |> flat)) Symtab.empty; fun deresolver name = (fst o the o AList.lookup (op =) ((fst o snd o the - o Symtab.lookup hs_program) ((mk_name_module o fst o dest_name) name))) name + o Symtab.lookup hs_program) ((mk_name_module o fst o Code_Namespace.dest_name) name))) name handle Option => error ("Unknown statement name: " ^ labelled_name name); in (deresolver, hs_program) end; fun serialize_haskell module_prefix string_classes { labelled_name, reserved_syms, includes, module_alias, class_syntax, tyco_syntax, const_syntax, program, - names, presentation_names } = + names } = let val reserved = fold (insert (op =) o fst) includes reserved_syms; val (deresolver, hs_program) = haskell_program_of_program labelled_name @@ -347,10 +365,10 @@ content, str "}" ]); - fun serialize_module1 (module_name', (deps, (stmts, _))) = + fun serialize_module (module_name', (deps, (stmts, _))) = let val stmt_names = map fst stmts; - val qualified = null presentation_names; + val qualified = true; val imports = subtract (op =) stmt_names deps |> distinct (op =) |> map_filter (try deresolver) @@ -363,18 +381,10 @@ val import_ps = map print_import_include includes @ map print_import_module imports val content = Pretty.chunks2 ((if null import_ps then [] else [Pretty.chunks import_ps]) @ map_filter - (fn (name, (_, SOME stmt)) => SOME (print_stmt qualified (name, stmt)) + (fn (name, (_, SOME stmt)) => SOME (markup_stmt name (print_stmt qualified (name, stmt))) | (_, (_, NONE)) => NONE) stmts ); in print_module module_name' content end; - fun serialize_module2 (_, (_, (stmts, _))) = Pretty.chunks2 (map_filter - (fn (name, (_, SOME stmt)) => if null presentation_names - orelse member (op =) presentation_names name - then SOME (print_stmt false (name, stmt)) - else NONE - | (_, (_, NONE)) => NONE) stmts); - val serialize_module = - if null presentation_names then serialize_module1 else pair "" o serialize_module2; fun write_module width (SOME destination) (modlname, content) = let val _ = File.check destination; @@ -386,13 +396,13 @@ val _ = File.mkdir_leaf (Path.dir pathname); in File.write pathname ("{-# OPTIONS_GHC -fglasgow-exts #-}\n\n" - ^ format false width content) + ^ format [] width content) end - | write_module width NONE (_, content) = writeln (format false width content); + | write_module width NONE (_, content) = writeln (format [] width content); in Code_Target.serialization (fn width => fn destination => K () o map (write_module width destination)) - (fn present => fn width => rpair [] o format present width o Pretty.chunks o map snd) + (fn present => fn width => rpair (fn _ => error "no deresolving") o format present width o Pretty.chunks o map snd) (map (uncurry print_module) includes @ map serialize_module (Symtab.dest hs_program)) end;
--- a/src/Tools/Code/code_ml.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_ml.ML Fri Sep 03 13:54:39 2010 +0200 @@ -759,10 +759,10 @@ fun modify_class stmts = single (SOME (ML_Class (the_single (map_filter (fn (name, Code_Thingol.Class (_, stmt)) => SOME (name, stmt) | _ => NONE) stmts)))) - fun modify_stmts ([stmt as (name, Code_Thingol.Fun _)]) = - [modify_fun stmt] + fun modify_stmts ([stmt as (name, stmt' as Code_Thingol.Fun _)]) = + if Code_Thingol.is_case stmt' then [] else [modify_fun stmt] | modify_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) = - modify_funs stmts + modify_funs (filter_out (Code_Thingol.is_case o snd) stmts) | modify_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) = modify_datatypes stmts | modify_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) = @@ -787,25 +787,21 @@ fun serialize_ml target print_module print_stmt with_signatures { labelled_name, reserved_syms, includes, module_alias, class_syntax, tyco_syntax, - const_syntax, program, names, presentation_names } = + const_syntax, program, names } = let val is_cons = Code_Thingol.is_cons program; - val is_presentation = not (null presentation_names); val { deresolver, hierarchical_program = ml_program } = ml_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program; + val print_stmt = print_stmt labelled_name tyco_syntax const_syntax + (make_vars reserved_syms) is_cons; fun print_node _ (_, Code_Namespace.Dummy) = NONE - | print_node prefix_fragments (_, Code_Namespace.Stmt stmt) = - if is_presentation andalso - (null o filter (member (op =) presentation_names) o stmt_names_of) stmt - then NONE - else SOME (print_stmt labelled_name tyco_syntax const_syntax (make_vars reserved_syms) - is_cons (deresolver prefix_fragments) stmt) + | print_node prefix_fragments (name, Code_Namespace.Stmt stmt) = + SOME (apsnd (markup_stmt name) (print_stmt (deresolver prefix_fragments) stmt)) | print_node prefix_fragments (name_fragment, Code_Namespace.Module (_, nodes)) = let val (decls, body) = print_nodes (prefix_fragments @ [name_fragment]) nodes; - val p = if is_presentation then Pretty.chunks2 body - else print_module name_fragment + val p = print_module name_fragment (if with_signatures then SOME decls else NONE) body; in SOME ([], p) end and print_nodes prefix_fragments nodes = (map_filter @@ -813,12 +809,11 @@ o rev o flat o Graph.strong_conn) nodes |> split_list |> (fn (decls, body) => (flat decls, body)) - val names' = map (try (deresolver [])) names; val p = Pretty.chunks2 (map snd includes @ snd (print_nodes [] ml_program)); - fun write width NONE = writeln o format false width - | write width (SOME p) = File.write p o format false width; + fun write width NONE = writeln o format [] width + | write width (SOME p) = File.write p o format [] width; in - Code_Target.serialization write (rpair names' ooo format) p + Code_Target.serialization write (rpair (try (deresolver [])) ooo format) p end; val serializer_sml : Code_Target.serializer =
--- a/src/Tools/Code/code_namespace.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_namespace.ML Fri Sep 03 13:54:39 2010 +0200 @@ -6,6 +6,7 @@ signature CODE_NAMESPACE = sig + val dest_name: string -> string * string datatype ('a, 'b) node = Dummy | Stmt of 'a @@ -23,7 +24,13 @@ structure Code_Namespace : CODE_NAMESPACE = struct -(* hierarchical program structure *) +(** splitting names in module and base part **) + +val dest_name = + apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode; + + +(** hierarchical program structure **) datatype ('a, 'b) node = Dummy @@ -46,10 +53,10 @@ of SOME name' => Long_Name.explode name' | NONE => map (fn name => fst (yield_singleton Name.variants name reserved)) (Long_Name.explode name); - val module_names = Graph.fold (insert (op =) o fst o Code_Printer.dest_name o fst) program []; + val module_names = Graph.fold (insert (op =) o fst o dest_name o fst) program []; val fragments_tab = fold (fn name => Symtab.update (name, alias_fragments name)) module_names Symtab.empty; - val dest_name = Code_Printer.dest_name #>> (the o Symtab.lookup fragments_tab); + val dest_name = dest_name #>> (the o Symtab.lookup fragments_tab); (* building empty module hierarchy *) val empty_module = (empty_data, Graph.empty);
--- a/src/Tools/Code/code_printer.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_printer.ML Fri Sep 03 13:54:39 2010 +0200 @@ -25,8 +25,8 @@ val semicolon: Pretty.T list -> Pretty.T val doublesemicolon: Pretty.T list -> Pretty.T val indent: int -> Pretty.T -> Pretty.T - val markup_stmt: string -> Pretty.T list -> Pretty.T - val format: bool -> int -> Pretty.T -> string + val markup_stmt: string -> Pretty.T -> Pretty.T + val format: string list -> int -> Pretty.T -> string val first_upper: string -> string val first_lower: string -> string @@ -70,7 +70,6 @@ val applify: string -> string -> ('a -> Pretty.T) -> fixity -> Pretty.T -> 'a list -> Pretty.T val tuplify: (fixity -> 'a -> Pretty.T) -> fixity -> 'a list -> Pretty.T option - type tyco_syntax type simple_const_syntax type complex_const_syntax @@ -93,10 +92,6 @@ val gen_print_bind: (thm option -> var_ctxt -> fixity -> iterm -> Pretty.T) -> thm option -> fixity -> iterm -> var_ctxt -> Pretty.T * var_ctxt - - val mk_name_module: Name.context -> string option -> (string -> string option) - -> 'a Graph.T -> string -> string - val dest_name: string -> string * string end; structure Code_Printer : CODE_PRINTER = @@ -110,7 +105,7 @@ | eqn_error NONE s = error s; val code_presentationN = "code_presentation"; -val _ = Output.add_mode code_presentationN; +val stmt_nameN = "stmt_name"; val _ = Markup.add_mode code_presentationN YXML.output_markup; @@ -132,21 +127,35 @@ fun doublesemicolon ps = Pretty.block [concat ps, str ";;"]; fun indent i = Print_Mode.setmp [] (Pretty.indent i); -val stmt_nameN = "stmt_name"; -fun markup_stmt name = Pretty.markup (code_presentationN, [(stmt_nameN, name)]); -fun filter_presentation selected (XML.Elem ((name, _), xs)) = - implode (map (filter_presentation (selected orelse name = code_presentationN)) xs) - | filter_presentation selected (XML.Text s) = +fun markup_stmt name = Print_Mode.setmp [code_presentationN] + (Pretty.mark (code_presentationN, [(stmt_nameN, name)])); + +fun filter_presentation presentation_names selected (XML.Elem ((name, attrs), xs)) = + implode (map (filter_presentation presentation_names + (selected orelse (name = code_presentationN + andalso member (op =) presentation_names (the (Properties.get attrs stmt_nameN))))) xs) + | filter_presentation presentation_names selected (XML.Text s) = if selected then s else ""; -fun format presentation_only width p = - if presentation_only then - Print_Mode.setmp [code_presentationN] (Pretty.string_of_margin width) p - |> YXML.parse_body - |> map (filter_presentation false) - |> implode - |> suffix "\n" - else Print_Mode.setmp [] (Pretty.string_of_margin width) p ^ "\n"; +fun maps_string s f [] = "" + | maps_string s f (x :: xs) = + let + val s1 = f x; + val s2 = maps_string s f xs; + in if s1 = "" then s2 + else if s2 = "" then s1 + else s1 ^ s ^ s2 + end; + +fun plain_text (XML.Elem (_, xs)) = implode (map plain_text xs) + | plain_text (XML.Text s) = s + +fun format presentation_names width = + Print_Mode.setmp [code_presentationN] (Pretty.string_of_margin width) + #> YXML.parse_body + #> (if null presentation_names then implode o map plain_text + else maps_string "\n\n" (filter_presentation presentation_names false)) + #> suffix "\n"; (** names and variable name contexts **) @@ -395,28 +404,4 @@ val parse_const_syntax = parse_syntax plain_const_syntax simple_const_syntax fst; - -(** module name spaces **) - -val dest_name = - apfst Long_Name.implode o split_last o fst o split_last o Long_Name.explode; - -fun mk_name_module reserved module_prefix module_alias program = - let - fun mk_alias name = case module_alias name - of SOME name' => name' - | NONE => name - |> Long_Name.explode - |> map (fn name => (the_single o fst) (Name.variants [name] reserved)) - |> Long_Name.implode; - fun mk_prefix name = case module_prefix - of SOME module_prefix => Long_Name.append module_prefix name - | NONE => name; - val tab = - Symtab.empty - |> Graph.fold ((fn name => Symtab.default (name, (mk_alias #> mk_prefix) name)) - o fst o dest_name o fst) - program - in the o Symtab.lookup tab end; - end; (*struct*)
--- a/src/Tools/Code/code_scala.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_scala.ML Fri Sep 03 13:54:39 2010 +0200 @@ -317,7 +317,8 @@ val implicits = filter (is_classinst o Graph.get_node program) (Graph.imm_succs program name); in union (op =) implicits end; - fun modify_stmt (_, Code_Thingol.Datatypecons _) = NONE + fun modify_stmt (_, Code_Thingol.Fun (_, (_, SOME _))) = NONE + | modify_stmt (_, Code_Thingol.Datatypecons _) = NONE | modify_stmt (_, Code_Thingol.Classrel _) = NONE | modify_stmt (_, Code_Thingol.Classparam _) = NONE | modify_stmt (_, stmt) = SOME stmt; @@ -329,7 +330,7 @@ fun serialize_scala { labelled_name, reserved_syms, includes, module_alias, class_syntax, tyco_syntax, const_syntax, program, - names, presentation_names } = + names } = let (* build program *) @@ -368,21 +369,15 @@ in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end; fun print_node _ (_, Code_Namespace.Dummy) = NONE | print_node prefix_fragments (name, Code_Namespace.Stmt stmt) = - if null presentation_names - orelse member (op =) presentation_names name - then SOME (print_stmt (deresolver prefix_fragments, deresolver []) (name, stmt)) - else NONE + SOME (markup_stmt name (print_stmt (deresolver prefix_fragments, deresolver []) (name, stmt))) | print_node prefix_fragments (name_fragment, Code_Namespace.Module (implicits, nodes)) = - if null presentation_names - then - let - val prefix_fragments' = prefix_fragments @ [name_fragment]; - in - Option.map (print_module name_fragment - (map_filter (print_implicit prefix_fragments') implicits)) - (print_nodes prefix_fragments' nodes) - end - else print_nodes [] nodes + let + val prefix_fragments' = prefix_fragments @ [name_fragment]; + in + Option.map (print_module name_fragment + (map_filter (print_implicit prefix_fragments') implicits)) + (print_nodes prefix_fragments' nodes) + end and print_nodes prefix_fragments nodes = let val ps = map_filter (fn name => print_node prefix_fragments (name, snd (Graph.get_node nodes name))) @@ -390,12 +385,11 @@ in if null ps then NONE else SOME (Pretty.chunks2 ps) end; (* serialization *) - val p_includes = if null presentation_names then map snd includes else []; - val p = Pretty.chunks2 (p_includes @ the_list (print_nodes [] sca_program)); - fun write width NONE = writeln o format false width - | write width (SOME p) = File.write p o format false width; + val p = Pretty.chunks2 (map snd includes @ the_list (print_nodes [] sca_program)); + fun write width NONE = writeln o format [] width + | write width (SOME p) = File.write p o format [] width; in - Code_Target.serialization write (rpair [] ooo format) p + Code_Target.serialization write (rpair (try (deresolver [])) ooo format) p end; val serializer : Code_Target.serializer =
--- a/src/Tools/Code/code_target.ML Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/Code/code_target.ML Fri Sep 03 13:54:39 2010 +0200 @@ -8,6 +8,7 @@ sig val cert_tyco: theory -> string -> string val read_tyco: theory -> string -> string + val read_const_exprs: theory -> string list -> string list val export_code_for: theory -> Path.T option -> string -> int option -> string -> Token.T list -> Code_Thingol.naming -> Code_Thingol.program -> string list -> unit @@ -42,7 +43,7 @@ type serialization val parse_args: 'a parser -> Token.T list -> 'a val serialization: (int -> Path.T option -> 'a -> unit) - -> (bool -> int -> 'a -> string * string option list) + -> (string list -> int -> 'a -> string * (string -> string option)) -> 'a -> serialization val set_default_code_width: int -> theory -> theory @@ -70,14 +71,11 @@ (** abstract nonsense **) datatype destination = Export of Path.T option | Produce | Present of string list; -type serialization = int -> destination -> (string * string option list) option; - -fun stmt_names_of_destination (Present stmt_names) = stmt_names - | stmt_names_of_destination _ = []; +type serialization = int -> destination -> (string * (string -> string option)) option; fun serialization output _ content width (Export some_path) = (output width some_path content; NONE) - | serialization _ string content width Produce = SOME (string false width content) - | serialization _ string content width (Present _) = SOME (string false width content); + | serialization _ string content width Produce = SOME (string [] width content) + | serialization _ string content width (Present stmt_names) = SOME (string stmt_names width content); fun export some_path f = (f (Export some_path); ()); fun produce f = the (f Produce); @@ -116,8 +114,7 @@ tyco_syntax: string -> Code_Printer.tyco_syntax option, const_syntax: string -> Code_Printer.activated_const_syntax option, program: Code_Thingol.program, - names: string list, - presentation_names: string list } + names: string list } -> serialization; datatype description = Fundamental of { serializer: serializer, @@ -308,7 +305,7 @@ fun invoke_serializer thy abortable serializer literals reserved abs_includes module_alias proto_class_syntax proto_instance_syntax proto_tyco_syntax proto_const_syntax - module_name args naming proto_program (names, presentation_names) = + module_name args naming proto_program names = let val (names_hidden, (class_syntax, tyco_syntax, const_syntax)) = activate_symbol_syntax thy literals naming @@ -325,19 +322,15 @@ tyco_syntax = Symtab.lookup tyco_syntax, const_syntax = Symtab.lookup const_syntax, program = program, - names = names, - presentation_names = presentation_names } + names = names } end; -fun mount_serializer thy target some_width raw_module_name args naming proto_program names destination = +fun mount_serializer thy target some_width module_name args naming proto_program names = let val (default_width, abortable, data, program) = activate_target_for thy target naming proto_program; val serializer = case the_description data of Fundamental seri => #serializer seri; - val presentation_names = stmt_names_of_destination destination; - val module_name = if null presentation_names - then raw_module_name else "Code"; val reserved = the_reserved data; fun select_include names_all (name, (content, cs)) = if null cs then SOME (name, content) @@ -354,7 +347,7 @@ in invoke_serializer thy abortable serializer literals reserved includes module_alias class instance tyco const module_name args - naming program (names, presentation_names) width destination + naming program names width end; fun assert_module_name "" = error ("Empty module name not allowed.") @@ -366,7 +359,9 @@ export some_path (mount_serializer thy target some_width some_module_name args naming program names); fun produce_code_for thy target some_width module_name args naming program names = - produce (mount_serializer thy target some_width (assert_module_name module_name) args naming program names); + let + val (s, deresolve) = produce (mount_serializer thy target some_width (assert_module_name module_name) args naming program names) + in (s, map deresolve names) end; fun present_code_for thy target some_width module_name args naming program (names, selects) = present selects (mount_serializer thy target some_width (assert_module_name module_name) args naming program names);
--- a/src/Tools/jEdit/src/jedit/document_view.scala Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/jEdit/src/jedit/document_view.scala Fri Sep 03 13:54:39 2010 +0200 @@ -24,7 +24,9 @@ object Document_View { - def choose_color(snapshot: Document.Snapshot, command: Command): Color = + /* physical rendering */ + + def status_color(snapshot: Document.Snapshot, command: Command): Color = { val state = snapshot.state(command) if (snapshot.is_outdated) new Color(240, 240, 240) @@ -38,6 +40,13 @@ } } + val message_markup: PartialFunction[Text.Info[Any], Color] = + { + case Text.Info(_, XML.Elem(Markup(Markup.WRITELN, _), _)) => new Color(220, 220, 220) + case Text.Info(_, XML.Elem(Markup(Markup.WARNING, _), _)) => new Color(255, 165, 0) + case Text.Info(_, XML.Elem(Markup(Markup.ERROR, _), _)) => new Color(255, 106, 106) + } + /* document view of text area */ @@ -163,18 +172,36 @@ Isabelle.swing_buffer_lock(model.buffer) { val snapshot = model.snapshot() val saved_color = gfx.getColor + val ascent = text_area.getPainter.getFontMetrics.getAscent try { for (i <- 0 until physical_lines.length) { if (physical_lines(i) != -1) { val line_range = proper_line_range(start(i), end(i)) + + // background color val cmds = snapshot.node.command_range(snapshot.revert(line_range)) - for ((command, command_start) <- cmds if !command.is_ignored) { + for { + (command, command_start) <- cmds if !command.is_ignored val range = line_range.restrict(snapshot.convert(command.range + command_start)) - val p = text_area.offsetToXY(range.start) - val q = text_area.offsetToXY(range.stop) - if (p != null && q != null) { - gfx.setColor(Document_View.choose_color(snapshot, command)) - gfx.fillRect(p.x, y + i * line_height, q.x - p.x, line_height) + r <- Isabelle.gfx_range(text_area, range) + } { + gfx.setColor(Document_View.status_color(snapshot, command)) + gfx.fillRect(r.x, y + i * line_height, r.length, line_height) + } + + // squiggly underline + for { + Text.Info(range, color) <- + snapshot.select_markup(line_range)(Document_View.message_markup)(null) + if color != null + r <- Isabelle.gfx_range(text_area, range) + } { + gfx.setColor(color) + val x0 = (r.x / 2) * 2 + val y0 = r.y + ascent + 1 + for (x1 <- Range(x0, x0 + r.length, 2)) { + val y1 = if (x1 % 4 < 2) y0 else y0 + 1 + gfx.drawLine(x1, y1, x1 + 1, y1) } } } @@ -266,7 +293,7 @@ val line2 = buffer.getLineOfOffset(snapshot.convert(start + command.length)) + 1 val y = line_to_y(line1) val height = HEIGHT * (line2 - line1) - gfx.setColor(Document_View.choose_color(snapshot, command)) + gfx.setColor(Document_View.status_color(snapshot, command)) gfx.fillRect(0, y, getWidth - 1, height) } }
--- a/src/Tools/jEdit/src/jedit/plugin.scala Fri Sep 03 13:54:04 2010 +0200 +++ b/src/Tools/jEdit/src/jedit/plugin.scala Fri Sep 03 13:54:39 2010 +0200 @@ -17,7 +17,7 @@ import org.gjt.sp.jedit.{jEdit, EBMessage, EBPlugin, Buffer, EditPane, ServiceManager, View} import org.gjt.sp.jedit.buffer.JEditBuffer -import org.gjt.sp.jedit.textarea.JEditTextArea +import org.gjt.sp.jedit.textarea.{JEditTextArea, TextArea} import org.gjt.sp.jedit.msg.{BufferUpdate, EditPaneUpdate, PropertiesChanged} import org.gjt.sp.jedit.gui.DockableWindowManager @@ -74,6 +74,19 @@ Int_Property("relative-font-size", 100)).toFloat / 100 + /* text area ranges */ + + case class Gfx_Range(val x: Int, val y: Int, val length: Int) + + def gfx_range(text_area: TextArea, range: Text.Range): Option[Gfx_Range] = + { + val p = text_area.offsetToXY(range.start) + val q = text_area.offsetToXY(range.stop) + if (p != null && q != null && p.y == q.y) Some(new Gfx_Range(p.x, p.y, q.x - p.x)) + else None + } + + /* tooltip markup */ def tooltip(text: String): String =