standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
--- a/src/Doc/Classes/Classes.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Classes/Classes.thy Tue Aug 13 16:25:47 2013 +0200
@@ -195,11 +195,11 @@
begin
definition %quote
- mult_prod_def: "p\<^isub>1 \<otimes> p\<^isub>2 = (fst p\<^isub>1 \<otimes> fst p\<^isub>2, snd p\<^isub>1 \<otimes> snd p\<^isub>2)"
+ mult_prod_def: "p\<^sub>1 \<otimes> p\<^sub>2 = (fst p\<^sub>1 \<otimes> fst p\<^sub>2, snd p\<^sub>1 \<otimes> snd p\<^sub>2)"
instance %quote proof
- fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "\<alpha>\<Colon>semigroup \<times> \<beta>\<Colon>semigroup"
- show "p\<^isub>1 \<otimes> p\<^isub>2 \<otimes> p\<^isub>3 = p\<^isub>1 \<otimes> (p\<^isub>2 \<otimes> p\<^isub>3)"
+ fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "\<alpha>\<Colon>semigroup \<times> \<beta>\<Colon>semigroup"
+ show "p\<^sub>1 \<otimes> p\<^sub>2 \<otimes> p\<^sub>3 = p\<^sub>1 \<otimes> (p\<^sub>2 \<otimes> p\<^sub>3)"
unfolding mult_prod_def by (simp add: assoc)
qed
--- a/src/Doc/Codegen/Foundations.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Codegen/Foundations.thy Tue Aug 13 16:25:47 2013 +0200
@@ -56,8 +56,8 @@
\noindent Central to code generation is the notion of \emph{code
equations}. A code equation as a first approximation is a theorem
- of the form @{text "f t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n \<equiv> t"} (an equation headed by a
- constant @{text f} with arguments @{text "t\<^isub>1 t\<^isub>2 \<dots> t\<^isub>n"} and right
+ of the form @{text "f t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n \<equiv> t"} (an equation headed by a
+ constant @{text f} with arguments @{text "t\<^sub>1 t\<^sub>2 \<dots> t\<^sub>n"} and right
hand side @{text t}).
\begin{itemize}
--- a/src/Doc/Codegen/Refinement.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Codegen/Refinement.thy Tue Aug 13 16:25:47 2013 +0200
@@ -116,8 +116,8 @@
to be.
The prerequisite for datatype constructors is only syntactical: a
- constructor must be of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^isub>1 \<dots> \<alpha>\<^isub>n"} where @{text
- "{\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n}"} is exactly the set of \emph{all} type variables in
+ constructor must be of type @{text "\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^sub>1 \<dots> \<alpha>\<^sub>n"} where @{text
+ "{\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n}"} is exactly the set of \emph{all} type variables in
@{text "\<tau>"}; then @{text "\<kappa>"} is its corresponding datatype. The
HOL datatype package by default registers any new datatype with its
constructors, but this may be changed using @{command_def
--- a/src/Doc/IsarImplementation/Logic.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarImplementation/Logic.thy Tue Aug 13 16:25:47 2013 +0200
@@ -38,18 +38,18 @@
\medskip A \emph{type class} is an abstract syntactic entity
declared in the theory context. The \emph{subclass relation} @{text
- "c\<^isub>1 \<subseteq> c\<^isub>2"} is specified by stating an acyclic
+ "c\<^sub>1 \<subseteq> c\<^sub>2"} is specified by stating an acyclic
generating relation; the transitive closure is maintained
internally. The resulting relation is an ordering: reflexive,
transitive, and antisymmetric.
- A \emph{sort} is a list of type classes written as @{text "s = {c\<^isub>1,
- \<dots>, c\<^isub>m}"}, it represents symbolic intersection. Notationally, the
+ A \emph{sort} is a list of type classes written as @{text "s = {c\<^sub>1,
+ \<dots>, c\<^sub>m}"}, it represents symbolic intersection. Notationally, the
curly braces are omitted for singleton intersections, i.e.\ any
class @{text "c"} may be read as a sort @{text "{c}"}. The ordering
on type classes is extended to sorts according to the meaning of
- intersections: @{text "{c\<^isub>1, \<dots> c\<^isub>m} \<subseteq> {d\<^isub>1, \<dots>, d\<^isub>n}"} iff @{text
- "\<forall>j. \<exists>i. c\<^isub>i \<subseteq> d\<^isub>j"}. The empty intersection @{text "{}"} refers to
+ intersections: @{text "{c\<^sub>1, \<dots> c\<^sub>m} \<subseteq> {d\<^sub>1, \<dots>, d\<^sub>n}"} iff @{text
+ "\<forall>j. \<exists>i. c\<^sub>i \<subseteq> d\<^sub>j"}. The empty intersection @{text "{}"} refers to
the universal sort, which is the largest element wrt.\ the sort
order. Thus @{text "{}"} represents the ``full sort'', not the
empty one! The intersection of all (finitely many) classes declared
@@ -57,10 +57,10 @@
\medskip A \emph{fixed type variable} is a pair of a basic name
(starting with a @{text "'"} character) and a sort constraint, e.g.\
- @{text "('a, s)"} which is usually printed as @{text "\<alpha>\<^isub>s"}.
+ @{text "('a, s)"} which is usually printed as @{text "\<alpha>\<^sub>s"}.
A \emph{schematic type variable} is a pair of an indexname and a
sort constraint, e.g.\ @{text "(('a, 0), s)"} which is usually
- printed as @{text "?\<alpha>\<^isub>s"}.
+ printed as @{text "?\<alpha>\<^sub>s"}.
Note that \emph{all} syntactic components contribute to the identity
of type variables: basic name, index, and sort constraint. The core
@@ -70,7 +70,7 @@
A \emph{type constructor} @{text "\<kappa>"} is a @{text "k"}-ary operator
on types declared in the theory. Type constructor application is
- written postfix as @{text "(\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>k)\<kappa>"}. For
+ written postfix as @{text "(\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>k)\<kappa>"}. For
@{text "k = 0"} the argument tuple is omitted, e.g.\ @{text "prop"}
instead of @{text "()prop"}. For @{text "k = 1"} the parentheses
are omitted, e.g.\ @{text "\<alpha> list"} instead of @{text "(\<alpha>)list"}.
@@ -79,7 +79,7 @@
\<beta>)fun"}.
The logical category \emph{type} is defined inductively over type
- variables and type constructors as follows: @{text "\<tau> = \<alpha>\<^isub>s | ?\<alpha>\<^isub>s |
+ variables and type constructors as follows: @{text "\<tau> = \<alpha>\<^sub>s | ?\<alpha>\<^sub>s |
(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"}.
A \emph{type abbreviation} is a syntactic definition @{text
@@ -89,27 +89,27 @@
logical core.
A \emph{type arity} declares the image behavior of a type
- constructor wrt.\ the algebra of sorts: @{text "\<kappa> :: (s\<^isub>1, \<dots>,
- s\<^isub>k)s"} means that @{text "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>k)\<kappa>"} is
- of sort @{text "s"} if every argument type @{text "\<tau>\<^isub>i"} is
- of sort @{text "s\<^isub>i"}. Arity declarations are implicitly
+ constructor wrt.\ the algebra of sorts: @{text "\<kappa> :: (s\<^sub>1, \<dots>,
+ s\<^sub>k)s"} means that @{text "(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>k)\<kappa>"} is
+ of sort @{text "s"} if every argument type @{text "\<tau>\<^sub>i"} is
+ of sort @{text "s\<^sub>i"}. Arity declarations are implicitly
completed, i.e.\ @{text "\<kappa> :: (\<^vec>s)c"} entails @{text "\<kappa> ::
(\<^vec>s)c'"} for any @{text "c' \<supseteq> c"}.
\medskip The sort algebra is always maintained as \emph{coregular},
which means that type arities are consistent with the subclass
relation: for any type constructor @{text "\<kappa>"}, and classes @{text
- "c\<^isub>1 \<subseteq> c\<^isub>2"}, and arities @{text "\<kappa> ::
- (\<^vec>s\<^isub>1)c\<^isub>1"} and @{text "\<kappa> ::
- (\<^vec>s\<^isub>2)c\<^isub>2"} holds @{text "\<^vec>s\<^isub>1 \<subseteq>
- \<^vec>s\<^isub>2"} component-wise.
+ "c\<^sub>1 \<subseteq> c\<^sub>2"}, and arities @{text "\<kappa> ::
+ (\<^vec>s\<^sub>1)c\<^sub>1"} and @{text "\<kappa> ::
+ (\<^vec>s\<^sub>2)c\<^sub>2"} holds @{text "\<^vec>s\<^sub>1 \<subseteq>
+ \<^vec>s\<^sub>2"} component-wise.
The key property of a coregular order-sorted algebra is that sort
constraints can be solved in a most general fashion: for each type
constructor @{text "\<kappa>"} and sort @{text "s"} there is a most general
- vector of argument sorts @{text "(s\<^isub>1, \<dots>, s\<^isub>k)"} such
- that a type scheme @{text "(\<alpha>\<^bsub>s\<^isub>1\<^esub>, \<dots>,
- \<alpha>\<^bsub>s\<^isub>k\<^esub>)\<kappa>"} is of sort @{text "s"}.
+ vector of argument sorts @{text "(s\<^sub>1, \<dots>, s\<^sub>k)"} such
+ that a type scheme @{text "(\<alpha>\<^bsub>s\<^sub>1\<^esub>, \<dots>,
+ \<alpha>\<^bsub>s\<^sub>k\<^esub>)\<kappa>"} is of sort @{text "s"}.
Consequently, type unification has most general solutions (modulo
equivalence of sorts), so type-inference produces primary types as
expected \cite{nipkow-prehofer}.
@@ -159,8 +159,8 @@
TVar}) in @{text "\<tau>"}; the type structure is traversed from left to
right.
- \item @{ML Sign.subsort}~@{text "thy (s\<^isub>1, s\<^isub>2)"}
- tests the subsort relation @{text "s\<^isub>1 \<subseteq> s\<^isub>2"}.
+ \item @{ML Sign.subsort}~@{text "thy (s\<^sub>1, s\<^sub>2)"}
+ tests the subsort relation @{text "s\<^sub>1 \<subseteq> s\<^sub>2"}.
\item @{ML Sign.of_sort}~@{text "thy (\<tau>, s)"} tests whether type
@{text "\<tau>"} is of sort @{text "s"}.
@@ -172,13 +172,13 @@
\item @{ML Sign.add_type_abbrev}~@{text "ctxt (\<kappa>, \<^vec>\<alpha>, \<tau>)"}
defines a new type abbreviation @{text "(\<^vec>\<alpha>)\<kappa> = \<tau>"}.
- \item @{ML Sign.primitive_class}~@{text "(c, [c\<^isub>1, \<dots>,
- c\<^isub>n])"} declares a new class @{text "c"}, together with class
- relations @{text "c \<subseteq> c\<^isub>i"}, for @{text "i = 1, \<dots>, n"}.
+ \item @{ML Sign.primitive_class}~@{text "(c, [c\<^sub>1, \<dots>,
+ c\<^sub>n])"} declares a new class @{text "c"}, together with class
+ relations @{text "c \<subseteq> c\<^sub>i"}, for @{text "i = 1, \<dots>, n"}.
- \item @{ML Sign.primitive_classrel}~@{text "(c\<^isub>1,
- c\<^isub>2)"} declares the class relation @{text "c\<^isub>1 \<subseteq>
- c\<^isub>2"}.
+ \item @{ML Sign.primitive_classrel}~@{text "(c\<^sub>1,
+ c\<^sub>2)"} declares the class relation @{text "c\<^sub>1 \<subseteq>
+ c\<^sub>2"}.
\item @{ML Sign.primitive_arity}~@{text "(\<kappa>, \<^vec>s, s)"} declares
the arity @{text "\<kappa> :: (\<^vec>s)s"}.
@@ -258,25 +258,25 @@
without any loose variables.
A \emph{fixed variable} is a pair of a basic name and a type, e.g.\
- @{text "(x, \<tau>)"} which is usually printed @{text "x\<^isub>\<tau>"} here. A
+ @{text "(x, \<tau>)"} which is usually printed @{text "x\<^sub>\<tau>"} here. A
\emph{schematic variable} is a pair of an indexname and a type,
e.g.\ @{text "((x, 0), \<tau>)"} which is likewise printed as @{text
- "?x\<^isub>\<tau>"}.
+ "?x\<^sub>\<tau>"}.
\medskip A \emph{constant} is a pair of a basic name and a type,
- e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text "c\<^isub>\<tau>"}
+ e.g.\ @{text "(c, \<tau>)"} which is usually printed as @{text "c\<^sub>\<tau>"}
here. Constants are declared in the context as polymorphic families
@{text "c :: \<sigma>"}, meaning that all substitution instances @{text
- "c\<^isub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
+ "c\<^sub>\<tau>"} for @{text "\<tau> = \<sigma>\<vartheta>"} are valid.
- The vector of \emph{type arguments} of constant @{text "c\<^isub>\<tau>"} wrt.\
+ The vector of \emph{type arguments} of constant @{text "c\<^sub>\<tau>"} wrt.\
the declaration @{text "c :: \<sigma>"} is defined as the codomain of the
- matcher @{text "\<vartheta> = {?\<alpha>\<^isub>1 \<mapsto> \<tau>\<^isub>1, \<dots>, ?\<alpha>\<^isub>n \<mapsto> \<tau>\<^isub>n}"} presented in
- canonical order @{text "(\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n)"}, corresponding to the
- left-to-right occurrences of the @{text "\<alpha>\<^isub>i"} in @{text "\<sigma>"}.
+ matcher @{text "\<vartheta> = {?\<alpha>\<^sub>1 \<mapsto> \<tau>\<^sub>1, \<dots>, ?\<alpha>\<^sub>n \<mapsto> \<tau>\<^sub>n}"} presented in
+ canonical order @{text "(\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n)"}, corresponding to the
+ left-to-right occurrences of the @{text "\<alpha>\<^sub>i"} in @{text "\<sigma>"}.
Within a given theory context, there is a one-to-one correspondence
- between any constant @{text "c\<^isub>\<tau>"} and the application @{text "c(\<tau>\<^isub>1,
- \<dots>, \<tau>\<^isub>n)"} of its type arguments. For example, with @{text "plus :: \<alpha>
+ between any constant @{text "c\<^sub>\<tau>"} and the application @{text "c(\<tau>\<^sub>1,
+ \<dots>, \<tau>\<^sub>n)"} of its type arguments. For example, with @{text "plus :: \<alpha>
\<Rightarrow> \<alpha> \<Rightarrow> \<alpha>"}, the instance @{text "plus\<^bsub>nat \<Rightarrow> nat \<Rightarrow> nat\<^esub>"} corresponds to
@{text "plus(nat)"}.
@@ -290,7 +290,7 @@
\medskip An \emph{atomic term} is either a variable or constant.
The logical category \emph{term} is defined inductively over atomic
terms, with abstraction and application as follows: @{text "t = b |
- x\<^isub>\<tau> | ?x\<^isub>\<tau> | c\<^isub>\<tau> | \<lambda>\<^isub>\<tau>. t | t\<^isub>1 t\<^isub>2"}. Parsing and printing takes care of
+ x\<^sub>\<tau> | ?x\<^sub>\<tau> | c\<^sub>\<tau> | \<lambda>\<^sub>\<tau>. t | t\<^sub>1 t\<^sub>2"}. Parsing and printing takes care of
converting between an external representation with named bound
variables. Subsequently, we shall use the latter notation instead
of internal de-Bruijn representation.
@@ -299,7 +299,7 @@
term according to the structure of atomic terms, abstractions, and
applicatins:
\[
- \infer{@{text "a\<^isub>\<tau> :: \<tau>"}}{}
+ \infer{@{text "a\<^sub>\<tau> :: \<tau>"}}{}
\qquad
\infer{@{text "(\<lambda>x\<^sub>\<tau>. t) :: \<tau> \<Rightarrow> \<sigma>"}}{@{text "t :: \<sigma>"}}
\qquad
@@ -315,7 +315,7 @@
The identity of atomic terms consists both of the name and the type
component. This means that different variables @{text
- "x\<^bsub>\<tau>\<^isub>1\<^esub>"} and @{text "x\<^bsub>\<tau>\<^isub>2\<^esub>"} may become the same after
+ "x\<^bsub>\<tau>\<^sub>1\<^esub>"} and @{text "x\<^bsub>\<tau>\<^sub>2\<^esub>"} may become the same after
type instantiation. Type-inference rejects variables of the same
name, but different types. In contrast, mixed instances of
polymorphic constants occur routinely.
@@ -329,11 +329,11 @@
pathological situation notoriously demands additional care.
\medskip A \emph{term abbreviation} is a syntactic definition @{text
- "c\<^isub>\<sigma> \<equiv> t"} of a closed term @{text "t"} of type @{text "\<sigma>"},
+ "c\<^sub>\<sigma> \<equiv> t"} of a closed term @{text "t"} of type @{text "\<sigma>"},
without any hidden polymorphism. A term abbreviation looks like a
constant in the syntax, but is expanded before entering the logical
core. Abbreviations are usually reverted when printing terms, using
- @{text "t \<rightarrow> c\<^isub>\<sigma>"} as rules for higher-order rewriting.
+ @{text "t \<rightarrow> c\<^sub>\<sigma>"} as rules for higher-order rewriting.
\medskip Canonical operations on @{text "\<lambda>"}-terms include @{text
"\<alpha>\<beta>\<eta>"}-conversion: @{text "\<alpha>"}-conversion refers to capture-free
@@ -428,7 +428,7 @@
introduces a new term abbreviation @{text "c \<equiv> t"}.
\item @{ML Sign.const_typargs}~@{text "thy (c, \<tau>)"} and @{ML
- Sign.const_instance}~@{text "thy (c, [\<tau>\<^isub>1, \<dots>, \<tau>\<^isub>n])"}
+ Sign.const_instance}~@{text "thy (c, [\<tau>\<^sub>1, \<dots>, \<tau>\<^sub>n])"}
convert between two representations of polymorphic constants: full
type instance vs.\ compact type arguments form.
@@ -497,7 +497,7 @@
The theory @{text "Pure"} contains constant declarations for the
primitive connectives @{text "\<And>"}, @{text "\<Longrightarrow>"}, and @{text "\<equiv>"} of
the logical framework, see \figref{fig:pure-connectives}. The
- derivability judgment @{text "A\<^isub>1, \<dots>, A\<^isub>n \<turnstile> B"} is
+ derivability judgment @{text "A\<^sub>1, \<dots>, A\<^sub>n \<turnstile> B"} is
defined inductively by the primitive inferences given in
\figref{fig:prim-rules}, with the global restriction that the
hypotheses must \emph{not} contain any schematic variables. The
@@ -564,7 +564,7 @@
"\<And>\<hyphen>intro"}) need not be recorded in the hypotheses, because
the simple syntactic types of Pure are always inhabitable.
``Assumptions'' @{text "x :: \<tau>"} for type-membership are only
- present as long as some @{text "x\<^isub>\<tau>"} occurs in the statement
+ present as long as some @{text "x\<^sub>\<tau>"} occurs in the statement
body.\footnote{This is the key difference to ``@{text "\<lambda>HOL"}'' in
the PTS framework \cite{Barendregt-Geuvers:2001}, where hypotheses
@{text "x : A"} are treated uniformly for propositions and types.}
@@ -625,7 +625,7 @@
"c((\<^vec>\<alpha>)\<kappa>) \<equiv> t"} for each type constructor @{text "\<kappa>"} (for
distinct variables @{text "\<^vec>\<alpha>"}). The RHS may mention
previously defined constants as above, or arbitrary constants @{text
- "d(\<alpha>\<^isub>i)"} for some @{text "\<alpha>\<^isub>i"} projected from @{text
+ "d(\<alpha>\<^sub>i)"} for some @{text "\<alpha>\<^sub>i"} projected from @{text
"\<^vec>\<alpha>"}. Thus overloaded definitions essentially work by
primitive recursion over the syntactic structure of a single type
argument. See also \cite[\S4.3]{Haftmann-Wenzel:2006:classes}.
@@ -738,10 +738,10 @@
variables are generalized simultaneously, specified by the given
basic names.
- \item @{ML Thm.instantiate}~@{text "(\<^vec>\<alpha>\<^isub>s,
- \<^vec>x\<^isub>\<tau>)"} corresponds to the @{text "instantiate"} rules
+ \item @{ML Thm.instantiate}~@{text "(\<^vec>\<alpha>\<^sub>s,
+ \<^vec>x\<^sub>\<tau>)"} corresponds to the @{text "instantiate"} rules
of \figref{fig:subst-rules}. Type variables are substituted before
- term variables. Note that the types in @{text "\<^vec>x\<^isub>\<tau>"}
+ term variables. Note that the types in @{text "\<^vec>x\<^sub>\<tau>"}
refer to the instantiated versions.
\item @{ML Thm.add_axiom}~@{text "ctxt (name, A)"} declares an
@@ -761,10 +761,10 @@
low-level representation in the axiom table may differ slightly from
the returned theorem.
- \item @{ML Theory.add_deps}~@{text "ctxt name c\<^isub>\<tau> \<^vec>d\<^isub>\<sigma>"}
+ \item @{ML Theory.add_deps}~@{text "ctxt name c\<^sub>\<tau> \<^vec>d\<^sub>\<sigma>"}
declares dependencies of a named specification for constant @{text
- "c\<^isub>\<tau>"}, relative to existing specifications for constants @{text
- "\<^vec>d\<^isub>\<sigma>"}.
+ "c\<^sub>\<tau>"}, relative to existing specifications for constants @{text
+ "\<^vec>d\<^sub>\<sigma>"}.
\end{description}
*}
@@ -1188,7 +1188,7 @@
\item @{text "rules\<^sub>1 RL rules\<^sub>2"} abbreviates @{text "rules\<^sub>1 RLN (1,
rules\<^sub>2)"}.
- \item @{text "[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule"} resolves @{text "rule\<^isub>i"}
+ \item @{text "[rule\<^sub>1, \<dots>, rule\<^sub>n] MRS rule"} resolves @{text "rule\<^sub>i"}
against premise @{text "i"} of @{text "rule"}, for @{text "i = n, \<dots>,
1"}. By working from right to left, newly emerging premises are
concatenated in the result, without interfering.
--- a/src/Doc/IsarImplementation/Prelim.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarImplementation/Prelim.thy Tue Aug 13 16:25:47 2013 +0200
@@ -844,7 +844,7 @@
\item @{ML Long_Name.qualifier}~@{text "name"} returns the qualifier
of a long name.
- \item @{ML Long_Name.append}~@{text "name\<^isub>1 name\<^isub>2"} appends two long
+ \item @{ML Long_Name.append}~@{text "name\<^sub>1 name\<^sub>2"} appends two long
names.
\item @{ML Long_Name.implode}~@{text "names"} and @{ML
@@ -989,7 +989,7 @@
\item Type @{ML_type Name_Space.T} represents name spaces.
\item @{ML Name_Space.empty}~@{text "kind"} and @{ML Name_Space.merge}~@{text
- "(space\<^isub>1, space\<^isub>2)"} are the canonical operations for
+ "(space\<^sub>1, space\<^sub>2)"} are the canonical operations for
maintaining name spaces according to theory data management
(\secref{sec:context-data}); @{text "kind"} is a formal comment
to characterize the purpose of a name space.
--- a/src/Doc/IsarImplementation/Proof.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarImplementation/Proof.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,7 +10,7 @@
Any variable that is not explicitly bound by @{text "\<lambda>"}-abstraction
is considered as ``free''. Logically, free variables act like
outermost universal quantification at the sequent level: @{text
- "A\<^isub>1(x), \<dots>, A\<^isub>n(x) \<turnstile> B(x)"} means that the result
+ "A\<^sub>1(x), \<dots>, A\<^sub>n(x) \<turnstile> B(x)"} means that the result
holds \emph{for all} values of @{text "x"}. Free variables for
terms (not types) can be fully internalized into the logic: @{text
"\<turnstile> B(x)"} and @{text "\<turnstile> \<And>x. B(x)"} are interchangeable, provided
@@ -35,26 +35,26 @@
depend on type variables, which means that type variables would have
to be declared first. For example, a raw type-theoretic framework
would demand the context to be constructed in stages as follows:
- @{text "\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^isub>\<alpha>)"}.
+ @{text "\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^sub>\<alpha>)"}.
We allow a slightly less formalistic mode of operation: term
variables @{text "x"} are fixed without specifying a type yet
(essentially \emph{all} potential occurrences of some instance
- @{text "x\<^isub>\<tau>"} are fixed); the first occurrence of @{text "x"}
+ @{text "x\<^sub>\<tau>"} are fixed); the first occurrence of @{text "x"}
within a specific term assigns its most general type, which is then
maintained consistently in the context. The above example becomes
- @{text "\<Gamma> = x: term, \<alpha>: type, A(x\<^isub>\<alpha>)"}, where type @{text
+ @{text "\<Gamma> = x: term, \<alpha>: type, A(x\<^sub>\<alpha>)"}, where type @{text
"\<alpha>"} is fixed \emph{after} term @{text "x"}, and the constraint
@{text "x :: \<alpha>"} is an implicit consequence of the occurrence of
- @{text "x\<^isub>\<alpha>"} in the subsequent proposition.
+ @{text "x\<^sub>\<alpha>"} in the subsequent proposition.
This twist of dependencies is also accommodated by the reverse
operation of exporting results from a context: a type variable
@{text "\<alpha>"} is considered fixed as long as it occurs in some fixed
term variable of the context. For example, exporting @{text "x:
- term, \<alpha>: type \<turnstile> x\<^isub>\<alpha> \<equiv> x\<^isub>\<alpha>"} produces in the first step @{text "x: term
- \<turnstile> x\<^isub>\<alpha> \<equiv> x\<^isub>\<alpha>"} for fixed @{text "\<alpha>"}, and only in the second step
- @{text "\<turnstile> ?x\<^isub>?\<^isub>\<alpha> \<equiv> ?x\<^isub>?\<^isub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
+ term, \<alpha>: type \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>"} produces in the first step @{text "x: term
+ \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>"} for fixed @{text "\<alpha>"}, and only in the second step
+ @{text "\<turnstile> ?x\<^sub>?\<^sub>\<alpha> \<equiv> ?x\<^sub>?\<^sub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
The following Isar source text illustrates this scenario.
*}
@@ -92,9 +92,9 @@
The @{text "focus"} operation provides a variant of @{text "import"}
for nested propositions (with explicit quantification): @{text
- "\<And>x\<^isub>1 \<dots> x\<^isub>n. B(x\<^isub>1, \<dots>, x\<^isub>n)"} is
- decomposed by inventing fixed variables @{text "x\<^isub>1, \<dots>,
- x\<^isub>n"} for the body.
+ "\<And>x\<^sub>1 \<dots> x\<^sub>n. B(x\<^sub>1, \<dots>, x\<^sub>n)"} is
+ decomposed by inventing fixed variables @{text "x\<^sub>1, \<dots>,
+ x\<^sub>n"} for the body.
*}
text %mlref {*
@@ -234,8 +234,8 @@
quantifiers stripped. For example, by assuming @{text "\<And>x :: \<alpha>. P
x"} we get @{text "\<And>x :: \<alpha>. P x \<turnstile> P ?x"} for schematic @{text "?x"}
of fixed type @{text "\<alpha>"}. Local derivations accumulate more and
- more explicit references to hypotheses: @{text "A\<^isub>1, \<dots>,
- A\<^isub>n \<turnstile> B"} where @{text "A\<^isub>1, \<dots>, A\<^isub>n"} needs to
+ more explicit references to hypotheses: @{text "A\<^sub>1, \<dots>,
+ A\<^sub>n \<turnstile> B"} where @{text "A\<^sub>1, \<dots>, A\<^sub>n"} needs to
be covered by the assumptions of the current context.
\medskip The @{text "add_assms"} operation augments the context by
--- a/src/Doc/IsarImplementation/Tactic.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarImplementation/Tactic.thy Tue Aug 13 16:25:47 2013 +0200
@@ -509,7 +509,7 @@
regarded as an atomic formula, to solve premise @{text "i"} of
@{text "thm\<^sub>2"}. Let @{text "thm\<^sub>1"} and @{text "thm\<^sub>2"} be @{text
"\<psi>"} and @{text "\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>"}. The unique @{text "s"} that
- unifies @{text "\<psi>"} and @{text "\<phi>\<^isub>i"} yields the theorem @{text "(\<phi>\<^sub>1 \<Longrightarrow>
+ unifies @{text "\<psi>"} and @{text "\<phi>\<^sub>i"} yields the theorem @{text "(\<phi>\<^sub>1 \<Longrightarrow>
\<dots> \<phi>\<^sub>i\<^sub>-\<^sub>1 \<Longrightarrow> \<phi>\<^sub>i\<^sub>+\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>)s"}. Multiple results are considered as
error (exception @{ML THM}).
--- a/src/Doc/IsarRef/First_Order_Logic.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/First_Order_Logic.thy Tue Aug 13 16:25:47 2013 +0200
@@ -183,15 +183,15 @@
axiomatization
disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) where
- disjI\<^isub>1 [intro]: "A \<Longrightarrow> A \<or> B" and
- disjI\<^isub>2 [intro]: "B \<Longrightarrow> A \<or> B" and
+ disjI\<^sub>1 [intro]: "A \<Longrightarrow> A \<or> B" and
+ disjI\<^sub>2 [intro]: "B \<Longrightarrow> A \<or> B" and
disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C"
axiomatization
conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) where
conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B" and
- conjD\<^isub>1: "A \<and> B \<Longrightarrow> A" and
- conjD\<^isub>2: "A \<and> B \<Longrightarrow> B"
+ conjD\<^sub>1: "A \<and> B \<Longrightarrow> A" and
+ conjD\<^sub>2: "A \<and> B \<Longrightarrow> B"
text {*
\noindent The conjunctive destructions have the disadvantage that
@@ -205,8 +205,8 @@
assumes "A \<and> B"
obtains A and B
proof
- from `A \<and> B` show A by (rule conjD\<^isub>1)
- from `A \<and> B` show B by (rule conjD\<^isub>2)
+ from `A \<and> B` show A by (rule conjD\<^sub>1)
+ from `A \<and> B` show B by (rule conjD\<^sub>2)
qed
text {*
@@ -378,8 +378,8 @@
@{text "impI: \<ASSUMES> A \<Longrightarrow> B \<SHOWS> A \<longrightarrow> B"} \\
@{text "impD: \<ASSUMES> A \<longrightarrow> B \<AND> A \<SHOWS> B"} \\[1ex]
- @{text "disjI\<^isub>1: \<ASSUMES> A \<SHOWS> A \<or> B"} \\
- @{text "disjI\<^isub>2: \<ASSUMES> B \<SHOWS> A \<or> B"} \\
+ @{text "disjI\<^sub>1: \<ASSUMES> A \<SHOWS> A \<or> B"} \\
+ @{text "disjI\<^sub>2: \<ASSUMES> B \<SHOWS> A \<or> B"} \\
@{text "disjE: \<ASSUMES> A \<or> B \<OBTAINS> A \<BBAR> B"} \\[1ex]
@{text "conjI: \<ASSUMES> A \<AND> B \<SHOWS> A \<and> B"} \\
--- a/src/Doc/IsarRef/Framework.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/Framework.thy Tue Aug 13 16:25:47 2013 +0200
@@ -270,8 +270,8 @@
:: \<alpha>. B(x)"} and @{text "A \<Longrightarrow> B"}. Primitive reasoning operates on
judgments of the form @{text "\<Gamma> \<turnstile> \<phi>"}, with standard introduction
and elimination rules for @{text "\<And>"} and @{text "\<Longrightarrow>"} that refer to
- fixed parameters @{text "x\<^isub>1, \<dots>, x\<^isub>m"} and hypotheses
- @{text "A\<^isub>1, \<dots>, A\<^isub>n"} from the context @{text "\<Gamma>"};
+ fixed parameters @{text "x\<^sub>1, \<dots>, x\<^sub>m"} and hypotheses
+ @{text "A\<^sub>1, \<dots>, A\<^sub>n"} from the context @{text "\<Gamma>"};
the corresponding proof terms are left implicit. The subsequent
inference rules define @{text "\<Gamma> \<turnstile> \<phi>"} inductively, relative to a
collection of axioms:
@@ -327,11 +327,11 @@
quantifiers are pulled in front of implications at each level of
nesting. This means that any Pure proposition may be presented as a
\emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the
- form @{text "\<And>x\<^isub>1 \<dots> x\<^isub>m. H\<^isub>1 \<Longrightarrow> \<dots> H\<^isub>n \<Longrightarrow>
+ form @{text "\<And>x\<^sub>1 \<dots> x\<^sub>m. H\<^sub>1 \<Longrightarrow> \<dots> H\<^sub>n \<Longrightarrow>
A"} for @{text "m, n \<ge> 0"}, and @{text "A"} atomic, and @{text
- "H\<^isub>1, \<dots>, H\<^isub>n"} being recursively of the same format.
+ "H\<^sub>1, \<dots>, H\<^sub>n"} being recursively of the same format.
Following the convention that outermost quantifiers are implicit,
- Horn clauses @{text "A\<^isub>1 \<Longrightarrow> \<dots> A\<^isub>n \<Longrightarrow> A"} are a special
+ Horn clauses @{text "A\<^sub>1 \<Longrightarrow> \<dots> A\<^sub>n \<Longrightarrow> A"} are a special
case of this.
For example, @{text "\<inter>"}-introduction rule encountered before is
@@ -348,9 +348,9 @@
@{text "InterI:"}~@{prop "(\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A) \<Longrightarrow> x \<in> \<Inter>\<A>"}
\]
- \medskip Goals are also represented as rules: @{text "A\<^isub>1 \<Longrightarrow>
- \<dots> A\<^isub>n \<Longrightarrow> C"} states that the sub-goals @{text "A\<^isub>1, \<dots>,
- A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the
+ \medskip Goals are also represented as rules: @{text "A\<^sub>1 \<Longrightarrow>
+ \<dots> A\<^sub>n \<Longrightarrow> C"} states that the sub-goals @{text "A\<^sub>1, \<dots>,
+ A\<^sub>n"} entail the result @{text "C"}; for @{text "n = 0"} the
goal is finished. To allow @{text "C"} being a rule statement
itself, we introduce the protective marker @{text "# :: prop \<Rightarrow>
prop"}, which is defined as identity and hidden from the user. We
@@ -369,7 +369,7 @@
sub-goal (replacing it by zero or more sub-goals), and @{inference
assumption}, for solving a sub-goal (finding a short-circuit with
local assumptions). Below @{text "\<^vec>x"} stands for @{text
- "x\<^isub>1, \<dots>, x\<^isub>n"} (@{text "n \<ge> 0"}).
+ "x\<^sub>1, \<dots>, x\<^sub>n"} (@{text "n \<ge> 0"}).
\[
\infer[(@{inference_def resolution})]
--- a/src/Doc/IsarRef/HOL_Specific.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/HOL_Specific.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1212,16 +1212,16 @@
\begin{matharray}{lll}
@{text "m"} & @{text "::"} &
- @{text "\<sigma>\<^isub>1 \<Rightarrow> \<dots> \<sigma>\<^isub>k \<Rightarrow> (\<^vec>\<alpha>\<^isub>n) t \<Rightarrow> (\<^vec>\<beta>\<^isub>n) t"} \\
+ @{text "\<sigma>\<^sub>1 \<Rightarrow> \<dots> \<sigma>\<^sub>k \<Rightarrow> (\<^vec>\<alpha>\<^sub>n) t \<Rightarrow> (\<^vec>\<beta>\<^sub>n) t"} \\
\end{matharray}
\noindent where @{text t} is the type constructor, @{text
- "\<^vec>\<alpha>\<^isub>n"} and @{text "\<^vec>\<beta>\<^isub>n"} are distinct
- type variables free in the local theory and @{text "\<sigma>\<^isub>1"},
- \ldots, @{text "\<sigma>\<^isub>k"} is a subsequence of @{text "\<alpha>\<^isub>1 \<Rightarrow>
- \<beta>\<^isub>1"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<alpha>\<^isub>1"}, \ldots,
- @{text "\<alpha>\<^isub>n \<Rightarrow> \<beta>\<^isub>n"}, @{text "\<beta>\<^isub>n \<Rightarrow>
- \<alpha>\<^isub>n"}.
+ "\<^vec>\<alpha>\<^sub>n"} and @{text "\<^vec>\<beta>\<^sub>n"} are distinct
+ type variables free in the local theory and @{text "\<sigma>\<^sub>1"},
+ \ldots, @{text "\<sigma>\<^sub>k"} is a subsequence of @{text "\<alpha>\<^sub>1 \<Rightarrow>
+ \<beta>\<^sub>1"}, @{text "\<beta>\<^sub>1 \<Rightarrow> \<alpha>\<^sub>1"}, \ldots,
+ @{text "\<alpha>\<^sub>n \<Rightarrow> \<beta>\<^sub>n"}, @{text "\<beta>\<^sub>n \<Rightarrow>
+ \<alpha>\<^sub>n"}.
\end{description}
*}
@@ -1441,7 +1441,7 @@
\begin{description}
- \item @{command "adhoc_overloading"}~@{text "c v\<^isub>1 ... v\<^isub>n"}
+ \item @{command "adhoc_overloading"}~@{text "c v\<^sub>1 ... v\<^sub>n"}
associates variants with an existing constant.
\item @{command "no_adhoc_overloading"} is similar to
@@ -1687,8 +1687,8 @@
\begin{description}
\item @{attribute (HOL) "coercion"}~@{text "f"} registers a new
- coercion function @{text "f :: \<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2"} where @{text "\<sigma>\<^isub>1"} and
- @{text "\<sigma>\<^isub>2"} are type constructors without arguments. Coercions are
+ coercion function @{text "f :: \<sigma>\<^sub>1 \<Rightarrow> \<sigma>\<^sub>2"} where @{text "\<sigma>\<^sub>1"} and
+ @{text "\<sigma>\<^sub>2"} are type constructors without arguments. Coercions are
composed by the inference algorithm if needed. Note that the type
inference algorithm is complete only if the registered coercions
form a lattice.
@@ -1699,11 +1699,11 @@
\begin{matharray}{lll}
@{text "map"} & @{text "::"} &
- @{text "f\<^isub>1 \<Rightarrow> \<dots> \<Rightarrow> f\<^isub>n \<Rightarrow> (\<alpha>\<^isub>1, \<dots>, \<alpha>\<^isub>n) t \<Rightarrow> (\<beta>\<^isub>1, \<dots>, \<beta>\<^isub>n) t"} \\
+ @{text "f\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> f\<^sub>n \<Rightarrow> (\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n) t \<Rightarrow> (\<beta>\<^sub>1, \<dots>, \<beta>\<^sub>n) t"} \\
\end{matharray}
- where @{text "t"} is a type constructor and @{text "f\<^isub>i"} is of type
- @{text "\<alpha>\<^isub>i \<Rightarrow> \<beta>\<^isub>i"} or @{text "\<beta>\<^isub>i \<Rightarrow> \<alpha>\<^isub>i"}. Registering a map function
+ where @{text "t"} is a type constructor and @{text "f\<^sub>i"} is of type
+ @{text "\<alpha>\<^sub>i \<Rightarrow> \<beta>\<^sub>i"} or @{text "\<beta>\<^sub>i \<Rightarrow> \<alpha>\<^sub>i"}. Registering a map function
overwrites any existing map function for this particular type
constructor.
@@ -1881,7 +1881,7 @@
(Ax - Bx) * (By - Cy) = (Ay - By) * (Bx - Cx)"
lemma collinear_inv_rotation:
- assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<twosuperior> + s\<twosuperior> = 1"
+ assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<^sup>2 + s\<^sup>2 = 1"
shows "collinear (Ax * c - Ay * s, Ay * c + Ax * s)
(Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)"
using assms by (algebra add: collinear.simps)
--- a/src/Doc/IsarRef/Misc.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/Misc.thy Tue Aug 13 16:25:47 2013 +0200
@@ -93,9 +93,9 @@
visualizes dependencies of facts, using Isabelle's graph browser
tool (see also \cite{isabelle-sys}).
- \item @{command "unused_thms"}~@{text "A\<^isub>1 \<dots> A\<^isub>m - B\<^isub>1 \<dots> B\<^isub>n"}
- displays all unused theorems in theories @{text "B\<^isub>1 \<dots> B\<^isub>n"}
- or their parents, but not in @{text "A\<^isub>1 \<dots> A\<^isub>m"} or their parents.
+ \item @{command "unused_thms"}~@{text "A\<^sub>1 \<dots> A\<^sub>m - B\<^sub>1 \<dots> B\<^sub>n"}
+ displays all unused theorems in theories @{text "B\<^sub>1 \<dots> B\<^sub>n"}
+ or their parents, but not in @{text "A\<^sub>1 \<dots> A\<^sub>m"} or their parents.
If @{text n} is @{text 0}, the end of the range of theories
defaults to the current theory. If no range is specified,
only the unused theorems in the current theory are displayed.
--- a/src/Doc/IsarRef/Outer_Syntax.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/Outer_Syntax.thy Tue Aug 13 16:25:47 2013 +0200
@@ -150,7 +150,7 @@
@{syntax_def verbatim} & = & @{verbatim "{*"} @{text "\<dots>"} @{verbatim "*"}@{verbatim "}"} \\[1ex]
@{text letter} & = & @{text "latin | "}@{verbatim "\\"}@{verbatim "<"}@{text latin}@{verbatim ">"}@{text " | "}@{verbatim "\\"}@{verbatim "<"}@{text "latin latin"}@{verbatim ">"}@{text " | greek |"} \\
- & & @{verbatim "\<^isub>"}@{text " | "}@{verbatim "\<^isup>"} \\
+ & & @{verbatim "\<^sub>"}@{text " | "}@{verbatim "\<^sup>"} \\
@{text quasiletter} & = & @{text "letter | digit | "}@{verbatim "_"}@{text " | "}@{verbatim "'"} \\
@{text latin} & = & @{verbatim a}@{text " | \<dots> | "}@{verbatim z}@{text " | "}@{verbatim A}@{text " | \<dots> | "}@{verbatim Z} \\
@{text digit} & = & @{verbatim "0"}@{text " | \<dots> | "}@{verbatim "9"} \\
--- a/src/Doc/IsarRef/Proof.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/Proof.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1224,8 +1224,8 @@
\item @{attribute case_names}~@{text "c\<^sub>1 \<dots> c\<^sub>k"} declares names for
the local contexts of premises of a theorem; @{text "c\<^sub>1, \<dots>, c\<^sub>k"}
refers to the \emph{prefix} of the list of premises. Each of the
- cases @{text "c\<^isub>i"} can be of the form @{text "c[h\<^isub>1 \<dots> h\<^isub>n]"} where
- the @{text "h\<^isub>1 \<dots> h\<^isub>n"} are the names of the hypotheses in case @{text "c\<^isub>i"}
+ cases @{text "c\<^sub>i"} can be of the form @{text "c[h\<^sub>1 \<dots> h\<^sub>n]"} where
+ the @{text "h\<^sub>1 \<dots> h\<^sub>n"} are the names of the hypotheses in case @{text "c\<^sub>i"}
from left to right.
\item @{attribute case_conclusion}~@{text "c d\<^sub>1 \<dots> d\<^sub>k"} declares
--- a/src/Doc/IsarRef/Spec.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/IsarRef/Spec.thy Tue Aug 13 16:25:47 2013 +0200
@@ -844,7 +844,7 @@
A weakend form of this is available through a further variant of
@{command instance}: @{command instance}~@{text "c\<^sub>1 \<subseteq> c\<^sub>2"} opens
- a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference
+ a proof that class @{text "c\<^sub>2"} implies @{text "c\<^sub>1"} without reference
to the underlying locales; this is useful if the properties to prove
the logical connection are not sufficent on the locale level but on
the theory level.
--- a/src/Doc/LaTeXsugar/Sugar.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/LaTeXsugar/Sugar.thy Tue Aug 13 16:25:47 2013 +0200
@@ -57,10 +57,10 @@
@{text"case"} are set in sans serif font to distinguish them from
other functions. This improves readability:
\begin{itemize}
-\item @{term"if b then e\<^isub>1 else e\<^isub>2"} instead of @{text"if b then e\<^isub>1 else e\<^isub>2"}.
-\item @{term"let x = e\<^isub>1 in e\<^isub>2"} instead of @{text"let x = e\<^isub>1 in e\<^isub>2"}.
-\item @{term"case x of True \<Rightarrow> e\<^isub>1 | False \<Rightarrow> e\<^isub>2"} instead of\\
- @{text"case x of True \<Rightarrow> e\<^isub>1 | False \<Rightarrow> e\<^isub>2"}.
+\item @{term"if b then e\<^sub>1 else e\<^sub>2"} instead of @{text"if b then e\<^sub>1 else e\<^sub>2"}.
+\item @{term"let x = e\<^sub>1 in e\<^sub>2"} instead of @{text"let x = e\<^sub>1 in e\<^sub>2"}.
+\item @{term"case x of True \<Rightarrow> e\<^sub>1 | False \<Rightarrow> e\<^sub>2"} instead of\\
+ @{text"case x of True \<Rightarrow> e\<^sub>1 | False \<Rightarrow> e\<^sub>2"}.
\end{itemize}
\subsection{Sets}
@@ -99,7 +99,7 @@
line. To avoid this, \texttt{OptionalSugar} contains syntax to group
@{text"@"}-terms to the left before printing, which leads to better
line breaking behaviour:
-@{term[display]"term\<^isub>0 @ term\<^isub>1 @ term\<^isub>2 @ term\<^isub>3 @ term\<^isub>4 @ term\<^isub>5 @ term\<^isub>6 @ term\<^isub>7 @ term\<^isub>8 @ term\<^isub>9 @ term\<^isub>1\<^isub>0"}
+@{term[display]"term\<^sub>0 @ term\<^sub>1 @ term\<^sub>2 @ term\<^sub>3 @ term\<^sub>4 @ term\<^sub>5 @ term\<^sub>6 @ term\<^sub>7 @ term\<^sub>8 @ term\<^sub>9 @ term\<^sub>1\<^sub>0"}
\end{itemize}
@@ -140,8 +140,8 @@
suppresses question marks; variables that end in digits,
e.g. @{text"x1"}, are still printed with a trailing @{text".0"},
e.g. @{text"x1.0"}, their internal index. This can be avoided by
-turning the last digit into a subscript: write \verb!x\<^isub>1! and
-obtain the much nicer @{text"x\<^isub>1"}. *}
+turning the last digit into a subscript: write \verb!x\<^sub>1! and
+obtain the much nicer @{text"x\<^sub>1"}. *}
(*<*)declare [[show_question_marks = false]](*>*)
--- a/src/Doc/Main/Main_Doc.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Main/Main_Doc.thy Tue Aug 13 16:25:47 2013 +0200
@@ -45,7 +45,7 @@
@{term"~(x = y)"} & @{term[source]"\<not> (x = y)"} & (\verb$~=$)\\
@{term[source]"P \<longleftrightarrow> Q"} & @{term"P \<longleftrightarrow> Q"} \\
@{term"If x y z"} & @{term[source]"If x y z"}\\
-@{term"Let e\<^isub>1 (%x. e\<^isub>2)"} & @{term[source]"Let e\<^isub>1 (\<lambda>x. e\<^isub>2)"}\\
+@{term"Let e\<^sub>1 (%x. e\<^sub>2)"} & @{term[source]"Let e\<^sub>1 (\<lambda>x. e\<^sub>2)"}\\
\end{supertabular}
@@ -131,7 +131,7 @@
\subsubsection*{Syntax}
\begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
-@{text"{x\<^isub>1,\<dots>,x\<^isub>n}"} & @{text"insert x\<^isub>1 (\<dots> (insert x\<^isub>n {})\<dots>)"}\\
+@{text"{x\<^sub>1,\<dots>,x\<^sub>n}"} & @{text"insert x\<^sub>1 (\<dots> (insert x\<^sub>n {})\<dots>)"}\\
@{term"x ~: A"} & @{term[source]"\<not>(x \<in> A)"}\\
@{term"A \<subseteq> B"} & @{term[source]"A \<le> B"}\\
@{term"A \<subset> B"} & @{term[source]"A < B"}\\
@@ -165,7 +165,7 @@
\begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
@{term"fun_upd f x y"} & @{term[source]"fun_upd f x y"}\\
-@{text"f(x\<^isub>1:=y\<^isub>1,\<dots>,x\<^isub>n:=y\<^isub>n)"} & @{text"f(x\<^isub>1:=y\<^isub>1)\<dots>(x\<^isub>n:=y\<^isub>n)"}\\
+@{text"f(x\<^sub>1:=y\<^sub>1,\<dots>,x\<^sub>n:=y\<^sub>n)"} & @{text"f(x\<^sub>1:=y\<^sub>1)\<dots>(x\<^sub>n:=y\<^sub>n)"}\\
\end{tabular}
@@ -545,7 +545,7 @@
\subsubsection*{Syntax}
\begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
-@{text"[x\<^isub>1,\<dots>,x\<^isub>n]"} & @{text"x\<^isub>1 # \<dots> # x\<^isub>n # []"}\\
+@{text"[x\<^sub>1,\<dots>,x\<^sub>n]"} & @{text"x\<^sub>1 # \<dots> # x\<^sub>n # []"}\\
@{term"[m..<n]"} & @{term[source]"upt m n"}\\
@{term"[i..j]"} & @{term[source]"upto i j"}\\
@{text"[e. x \<leftarrow> xs]"} & @{term"map (%x. e) xs"}\\
@@ -555,8 +555,8 @@
\end{supertabular}
\medskip
-List comprehension: @{text"[e. q\<^isub>1, \<dots>, q\<^isub>n]"} where each
-qualifier @{text q\<^isub>i} is either a generator \mbox{@{text"pat \<leftarrow> e"}} or a
+List comprehension: @{text"[e. q\<^sub>1, \<dots>, q\<^sub>n]"} where each
+qualifier @{text q\<^sub>i} is either a generator \mbox{@{text"pat \<leftarrow> e"}} or a
guard, i.e.\ boolean expression.
\section*{Map}
@@ -581,8 +581,8 @@
\begin{tabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
@{term"Map.empty"} & @{term"\<lambda>x. None"}\\
@{term"m(x:=Some y)"} & @{term[source]"m(x:=Some y)"}\\
-@{text"m(x\<^isub>1\<mapsto>y\<^isub>1,\<dots>,x\<^isub>n\<mapsto>y\<^isub>n)"} & @{text[source]"m(x\<^isub>1\<mapsto>y\<^isub>1)\<dots>(x\<^isub>n\<mapsto>y\<^isub>n)"}\\
-@{text"[x\<^isub>1\<mapsto>y\<^isub>1,\<dots>,x\<^isub>n\<mapsto>y\<^isub>n]"} & @{text[source]"Map.empty(x\<^isub>1\<mapsto>y\<^isub>1,\<dots>,x\<^isub>n\<mapsto>y\<^isub>n)"}\\
+@{text"m(x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n)"} & @{text[source]"m(x\<^sub>1\<mapsto>y\<^sub>1)\<dots>(x\<^sub>n\<mapsto>y\<^sub>n)"}\\
+@{text"[x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n]"} & @{text[source]"Map.empty(x\<^sub>1\<mapsto>y\<^sub>1,\<dots>,x\<^sub>n\<mapsto>y\<^sub>n)"}\\
@{term"map_upds m xs ys"} & @{term[source]"map_upds m xs ys"}\\
\end{tabular}
--- a/src/Doc/ProgProve/Basics.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/ProgProve/Basics.thy Tue Aug 13 16:25:47 2013 +0200
@@ -35,8 +35,8 @@
\concept{Terms} are formed as in functional programming by
applying functions to arguments. If @{text f} is a function of type
-@{text"\<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2"} and @{text t} is a term of type
-@{text"\<tau>\<^isub>1"} then @{term"f t"} is a term of type @{text"\<tau>\<^isub>2"}. We write @{text "t :: \<tau>"} to mean that term @{text t} has type @{text \<tau>}.
+@{text"\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2"} and @{text t} is a term of type
+@{text"\<tau>\<^sub>1"} then @{term"f t"} is a term of type @{text"\<tau>\<^sub>2"}. We write @{text "t :: \<tau>"} to mean that term @{text t} has type @{text \<tau>}.
\begin{warn}
There are many predefined infix symbols like @{text "+"} and @{text"\<le>"}.
@@ -47,9 +47,9 @@
HOL also supports some basic constructs from functional programming:
\begin{quote}
-@{text "(if b then t\<^isub>1 else t\<^isub>2)"}\\
+@{text "(if b then t\<^sub>1 else t\<^sub>2)"}\\
@{text "(let x = t in u)"}\\
-@{text "(case t of pat\<^isub>1 \<Rightarrow> t\<^isub>1 | \<dots> | pat\<^isub>n \<Rightarrow> t\<^isub>n)"}
+@{text "(case t of pat\<^sub>1 \<Rightarrow> t\<^sub>1 | \<dots> | pat\<^sub>n \<Rightarrow> t\<^sub>n)"}
\end{quote}
\begin{warn}
The above three constructs must always be enclosed in parentheses
@@ -86,11 +86,11 @@
\begin{warn}
Right-arrows of all kinds always associate to the right. In particular,
the formula
-@{text"A\<^isub>1 \<Longrightarrow> A\<^isub>2 \<Longrightarrow> A\<^isub>3"} means @{text "A\<^isub>1 \<Longrightarrow> (A\<^isub>2 \<Longrightarrow> A\<^isub>3)"}.
-The (Isabelle specific) notation \mbox{@{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"}}
-is short for the iterated implication \mbox{@{text"A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> A"}}.
+@{text"A\<^sub>1 \<Longrightarrow> A\<^sub>2 \<Longrightarrow> A\<^sub>3"} means @{text "A\<^sub>1 \<Longrightarrow> (A\<^sub>2 \<Longrightarrow> A\<^sub>3)"}.
+The (Isabelle specific) notation \mbox{@{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"}}
+is short for the iterated implication \mbox{@{text"A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> A"}}.
Sometimes we also employ inference rule notation:
-\inferrule{\mbox{@{text "A\<^isub>1"}}\\ \mbox{@{text "\<dots>"}}\\ \mbox{@{text "A\<^isub>n"}}}
+\inferrule{\mbox{@{text "A\<^sub>1"}}\\ \mbox{@{text "\<dots>"}}\\ \mbox{@{text "A\<^sub>n"}}}
{\mbox{@{text A}}}
\end{warn}
@@ -104,13 +104,13 @@
The general format of a theory @{text T} is
\begin{quote}
\isacom{theory} @{text T}\\
-\isacom{imports} @{text "T\<^isub>1 \<dots> T\<^isub>n"}\\
+\isacom{imports} @{text "T\<^sub>1 \<dots> T\<^sub>n"}\\
\isacom{begin}\\
\emph{definitions, theorems and proofs}\\
\isacom{end}
\end{quote}
-where @{text "T\<^isub>1 \<dots> T\<^isub>n"} are the names of existing
-theories that @{text T} is based on. The @{text "T\<^isub>i"} are the
+where @{text "T\<^sub>1 \<dots> T\<^sub>n"} are the names of existing
+theories that @{text T} is based on. The @{text "T\<^sub>i"} are the
direct \concept{parent theories} of @{text T}.
Everything defined in the parent theories (and their parents, recursively) is
automatically visible. Each theory @{text T} must
--- a/src/Doc/ProgProve/Bool_nat_list.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/ProgProve/Bool_nat_list.thy Tue Aug 13 16:25:47 2013 +0200
@@ -385,7 +385,7 @@
\begin{itemize}
\item @{text "[]"} is @{const Nil},
\item @{term"x # xs"} is @{term"Cons x xs"},
-\item @{text"[x\<^isub>1, \<dots>, x\<^isub>n]"} is @{text"x\<^isub>1 # \<dots> # x\<^isub>n # []"}, and
+\item @{text"[x\<^sub>1, \<dots>, x\<^sub>n]"} is @{text"x\<^sub>1 # \<dots> # x\<^sub>n # []"}, and
\item @{term "xs @ ys"} is @{term"app xs ys"}.
\end{itemize}
There is also a large library of predefined functions.
--- a/src/Doc/ProgProve/Isar.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/ProgProve/Isar.thy Tue Aug 13 16:25:47 2013 +0200
@@ -443,8 +443,8 @@
@{text"\<longleftrightarrow>"}:
\end{isamarkuptext}%
*}
-(*<*)lemma "formula\<^isub>1 \<longleftrightarrow> formula\<^isub>2" proof-(*>*)
-show "formula\<^isub>1 \<longleftrightarrow> formula\<^isub>2" (is "?L \<longleftrightarrow> ?R")
+(*<*)lemma "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" proof-(*>*)
+show "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" (is "?L \<longleftrightarrow> ?R")
proof
assume "?L"
txt_raw{*\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}*}
@@ -455,7 +455,7 @@
show "?L" (*<*)sorry(*>*) txt_raw{*\ $\dots$\\*}
qed(*<*)qed(*>*)
-text{* Instead of duplicating @{text"formula\<^isub>i"} in the text, we introduce
+text{* Instead of duplicating @{text"formula\<^sub>i"} in the text, we introduce
the two abbreviations @{text"?L"} and @{text"?R"} by pattern matching.
Pattern matching works wherever a formula is stated, in particular
with \isacom{have} and \isacom{lemma}.
@@ -513,11 +513,11 @@
\isa{%
*}
(*<*)lemma "P" proof-(*>*)
-have "P\<^isub>1" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
-moreover have "P\<^isub>2" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
+have "P\<^sub>1" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
+moreover have "P\<^sub>2" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
moreover
txt_raw{*\\$\vdots$\\\hspace{-1.4ex}*}(*<*)have "True" ..(*>*)
-moreover have "P\<^isub>n" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
+moreover have "P\<^sub>n" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
ultimately have "P" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
(*<*)oops(*>*)
@@ -529,11 +529,11 @@
\isa{%
*}
(*<*)lemma "P" proof-(*>*)
-have lab\<^isub>1: "P\<^isub>1" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
-have lab\<^isub>2: "P\<^isub>2" (*<*)sorry(*>*)txt_raw{*\ $\dots$*}
+have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
+have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)txt_raw{*\ $\dots$*}
txt_raw{*\\$\vdots$\\\hspace{-1.4ex}*}
-have lab\<^isub>n: "P\<^isub>n" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
-from lab\<^isub>1 lab\<^isub>2 txt_raw{*\ $\dots$\\*}
+have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
+from lab\<^sub>1 lab\<^sub>2 txt_raw{*\ $\dots$\\*}
have "P" (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
(*<*)oops(*>*)
@@ -551,14 +551,14 @@
has its own assumptions and is generalized over its locally fixed
variables at the end. This is what a \concept{raw proof block} does:
\begin{quote}
-@{text"{"} \isacom{fix} @{text"x\<^isub>1 \<dots> x\<^isub>n"}\\
-\mbox{}\ \ \ \isacom{assume} @{text"A\<^isub>1 \<dots> A\<^isub>m"}\\
+@{text"{"} \isacom{fix} @{text"x\<^sub>1 \<dots> x\<^sub>n"}\\
+\mbox{}\ \ \ \isacom{assume} @{text"A\<^sub>1 \<dots> A\<^sub>m"}\\
\mbox{}\ \ \ $\vdots$\\
\mbox{}\ \ \ \isacom{have} @{text"B"}\\
@{text"}"}
\end{quote}
-proves @{text"\<lbrakk> A\<^isub>1; \<dots> ; A\<^isub>m \<rbrakk> \<Longrightarrow> B"}
-where all @{text"x\<^isub>i"} have been replaced by unknowns @{text"?x\<^isub>i"}.
+proves @{text"\<lbrakk> A\<^sub>1; \<dots> ; A\<^sub>m \<rbrakk> \<Longrightarrow> B"}
+where all @{text"x\<^sub>i"} have been replaced by unknowns @{text"?x\<^sub>i"}.
\begin{warn}
The conclusion of a raw proof block is \emph{not} indicated by \isacom{show}
but is simply the final \isacom{have}.
@@ -642,15 +642,15 @@
This proof pattern works for any term @{text t} whose type is a datatype.
The goal has to be proved for each constructor @{text C}:
\begin{quote}
-\isacom{fix} \ @{text"x\<^isub>1 \<dots> x\<^isub>n"} \isacom{assume} @{text"\"t = C x\<^isub>1 \<dots> x\<^isub>n\""}
+\isacom{fix} \ @{text"x\<^sub>1 \<dots> x\<^sub>n"} \isacom{assume} @{text"\"t = C x\<^sub>1 \<dots> x\<^sub>n\""}
\end{quote}
Each case can be written in a more compact form by means of the \isacom{case}
command:
\begin{quote}
-\isacom{case} @{text "(C x\<^isub>1 \<dots> x\<^isub>n)"}
+\isacom{case} @{text "(C x\<^sub>1 \<dots> x\<^sub>n)"}
\end{quote}
This is equivalent to the explicit \isacom{fix}-\isacom{assume} line
-but also gives the assumption @{text"\"t = C x\<^isub>1 \<dots> x\<^isub>n\""} a name: @{text C},
+but also gives the assumption @{text"\"t = C x\<^sub>1 \<dots> x\<^sub>n\""} a name: @{text C},
like the constructor.
Here is the \isacom{case} version of the proof above:
*}
@@ -782,16 +782,16 @@
(here @{text"A(Suc n)"}).
Induction works for any datatype.
-Proving a goal @{text"\<lbrakk> A\<^isub>1(x); \<dots>; A\<^isub>k(x) \<rbrakk> \<Longrightarrow> P(x)"}
+Proving a goal @{text"\<lbrakk> A\<^sub>1(x); \<dots>; A\<^sub>k(x) \<rbrakk> \<Longrightarrow> P(x)"}
by induction on @{text x} generates a proof obligation for each constructor
-@{text C} of the datatype. The command @{text"case (C x\<^isub>1 \<dots> x\<^isub>n)"}
+@{text C} of the datatype. The command @{text"case (C x\<^sub>1 \<dots> x\<^sub>n)"}
performs the following steps:
\begin{enumerate}
-\item \isacom{fix} @{text"x\<^isub>1 \<dots> x\<^isub>n"}
+\item \isacom{fix} @{text"x\<^sub>1 \<dots> x\<^sub>n"}
\item \isacom{assume} the induction hypotheses (calling them @{text C.IH})
- and the premises \mbox{@{text"A\<^isub>i(C x\<^isub>1 \<dots> x\<^isub>n)"}} (calling them @{text"C.prems"})
+ and the premises \mbox{@{text"A\<^sub>i(C x\<^sub>1 \<dots> x\<^sub>n)"}} (calling them @{text"C.prems"})
and calling the whole list @{text C}
-\item \isacom{let} @{text"?case = \"P(C x\<^isub>1 \<dots> x\<^isub>n)\""}
+\item \isacom{let} @{text"?case = \"P(C x\<^sub>1 \<dots> x\<^sub>n)\""}
\end{enumerate}
\subsection{Rule Induction}
@@ -889,34 +889,34 @@
\indent
In general, let @{text I} be a (for simplicity unary) inductively defined
predicate and let the rules in the definition of @{text I}
-be called @{text "rule\<^isub>1"}, \dots, @{text "rule\<^isub>n"}. A proof by rule
+be called @{text "rule\<^sub>1"}, \dots, @{text "rule\<^sub>n"}. A proof by rule
induction follows this pattern:
*}
(*<*)
-inductive I where rule\<^isub>1: "I()" | rule\<^isub>2: "I()" | rule\<^isub>n: "I()"
+inductive I where rule\<^sub>1: "I()" | rule\<^sub>2: "I()" | rule\<^sub>n: "I()"
lemma "I x \<Longrightarrow> P x" proof-(*>*)
show "I x \<Longrightarrow> P x"
proof(induction rule: I.induct)
- case rule\<^isub>1
+ case rule\<^sub>1
txt_raw{*\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*}
show ?case (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
next
txt_raw{*\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*}
(*<*)
- case rule\<^isub>2
+ case rule\<^sub>2
show ?case sorry
(*>*)
next
- case rule\<^isub>n
+ case rule\<^sub>n
txt_raw{*\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}*}
show ?case (*<*)sorry(*>*)txt_raw{*\ $\dots$\\*}
qed(*<*)qed(*>*)
text{*
One can provide explicit variable names by writing
-\isacom{case}~@{text"(rule\<^isub>i x\<^isub>1 \<dots> x\<^isub>k)"}, thus renaming the first @{text k}
-free variables in rule @{text i} to @{text"x\<^isub>1 \<dots> x\<^isub>k"},
+\isacom{case}~@{text"(rule\<^sub>i x\<^sub>1 \<dots> x\<^sub>k)"}, thus renaming the first @{text k}
+free variables in rule @{text i} to @{text"x\<^sub>1 \<dots> x\<^sub>k"},
going through rule @{text i} from left to right.
\subsection{Assumption Naming}
@@ -930,8 +930,8 @@
induction rule. For rule inductions these are the hypotheses of rule
@{text name}, for structural inductions these are empty.
\item[@{text name.prems}] contains the (suitably instantiated) premises
-of the statement being proved, i.e. the @{text A\<^isub>i} when
-proving @{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"}.
+of the statement being proved, i.e. the @{text A\<^sub>i} when
+proving @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"}.
\end{description}
\begin{warn}
Proof method @{text induct} differs from @{text induction}
--- a/src/Doc/ProgProve/Logic.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/ProgProve/Logic.thy Tue Aug 13 16:25:47 2013 +0200
@@ -72,8 +72,8 @@
theorems and proof states, separating assumptions from conclusions.
The implication @{text"\<longrightarrow>"} is part of the logic HOL and can occur inside the
formulas that make up the assumptions and conclusion.
-Theorems should be of the form @{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"},
-not @{text"A\<^isub>1 \<and> \<dots> \<and> A\<^isub>n \<longrightarrow> A"}. Both are logically equivalent
+Theorems should be of the form @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"},
+not @{text"A\<^sub>1 \<and> \<dots> \<and> A\<^sub>n \<longrightarrow> A"}. Both are logically equivalent
but the first one works better when using the theorem in further proofs.
\end{warn}
@@ -83,7 +83,7 @@
Sets of elements of type @{typ 'a} have type @{typ"'a set"}.
They can be finite or infinite. Sets come with the usual notation:
\begin{itemize}
-\item @{term"{}"},\quad @{text"{e\<^isub>1,\<dots>,e\<^isub>n}"}
+\item @{term"{}"},\quad @{text"{e\<^sub>1,\<dots>,e\<^sub>n}"}
\item @{prop"e \<in> A"},\quad @{prop"A \<subseteq> B"}
\item @{term"A \<union> B"},\quad @{term"A \<inter> B"},\quad @{term"A - B"},\quad @{term"-A"}
\end{itemize}
@@ -319,9 +319,9 @@
two formulas @{text"a=b"} and @{text False}, yielding the rule
@{thm[display,mode=Rule]conjI[of "a=b" False]}
-In general, @{text"th[of string\<^isub>1 \<dots> string\<^isub>n]"} instantiates
+In general, @{text"th[of string\<^sub>1 \<dots> string\<^sub>n]"} instantiates
the unknowns in the theorem @{text th} from left to right with the terms
-@{text string\<^isub>1} to @{text string\<^isub>n}.
+@{text string\<^sub>1} to @{text string\<^sub>n}.
\item By unification. \concept{Unification} is the process of making two
terms syntactically equal by suitable instantiations of unknowns. For example,
@@ -346,13 +346,13 @@
@{text"1. \<dots> \<Longrightarrow> A"}\\
@{text"2. \<dots> \<Longrightarrow> B"}
\end{quote}
-In general, the application of a rule @{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"}
+In general, the application of a rule @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"}
to a subgoal \mbox{@{text"\<dots> \<Longrightarrow> C"}} proceeds in two steps:
\begin{enumerate}
\item
Unify @{text A} and @{text C}, thus instantiating the unknowns in the rule.
\item
-Replace the subgoal @{text C} with @{text n} new subgoals @{text"A\<^isub>1"} to @{text"A\<^isub>n"}.
+Replace the subgoal @{text C} with @{text n} new subgoals @{text"A\<^sub>1"} to @{text"A\<^sub>n"}.
\end{enumerate}
This is the command to apply rule @{text xyz}:
\begin{quote}
@@ -430,10 +430,10 @@
premises.
\end{warn}
-In general @{text r} can be of the form @{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"}
-and there can be multiple argument theorems @{text r\<^isub>1} to @{text r\<^isub>m}
-(with @{text"m \<le> n"}), in which case @{text "r[OF r\<^isub>1 \<dots> r\<^isub>m]"} is obtained
-by unifying and thus proving @{text "A\<^isub>i"} with @{text "r\<^isub>i"}, @{text"i = 1\<dots>m"}.
+In general @{text r} can be of the form @{text"\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A"}
+and there can be multiple argument theorems @{text r\<^sub>1} to @{text r\<^sub>m}
+(with @{text"m \<le> n"}), in which case @{text "r[OF r\<^sub>1 \<dots> r\<^sub>m]"} is obtained
+by unifying and thus proving @{text "A\<^sub>i"} with @{text "r\<^sub>i"}, @{text"i = 1\<dots>m"}.
Here is an example, where @{thm[source]refl} is the theorem
@{thm[show_question_marks] refl}:
*}
@@ -739,7 +739,7 @@
\end{quote}
followed by a sequence of (possibly named) rules of the form
\begin{quote}
-@{text "\<lbrakk> I a\<^isub>1; \<dots>; I a\<^isub>n \<rbrakk> \<Longrightarrow> I a"}
+@{text "\<lbrakk> I a\<^sub>1; \<dots>; I a\<^sub>n \<rbrakk> \<Longrightarrow> I a"}
\end{quote}
separated by @{text"|"}. As usual, @{text n} can be 0.
The corresponding rule induction principle
@@ -755,7 +755,7 @@
Proving @{prop "I x \<Longrightarrow> P x"} by rule induction means proving
for every rule of @{text I} that @{text P} is invariant:
\begin{quote}
-@{text "\<lbrakk> I a\<^isub>1; P a\<^isub>1; \<dots>; I a\<^isub>n; P a\<^isub>n \<rbrakk> \<Longrightarrow> P a"}
+@{text "\<lbrakk> I a\<^sub>1; P a\<^sub>1; \<dots>; I a\<^sub>n; P a\<^sub>n \<rbrakk> \<Longrightarrow> P a"}
\end{quote}
The above format for inductive definitions is simplified in a number of
--- a/src/Doc/ProgProve/Types_and_funs.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/ProgProve/Types_and_funs.thy Tue Aug 13 16:25:47 2013 +0200
@@ -21,14 +21,14 @@
The general form of a datatype definition looks like this:
\begin{quote}
\begin{tabular}{@ {}rclcll}
-\isacom{datatype} @{text "('a\<^isub>1,\<dots>,'a\<^isub>n)t"}
+\isacom{datatype} @{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"}
& = & $C_1 \ @{text"\""}\tau_{1,1}@{text"\""} \dots @{text"\""}\tau_{1,n_1}@{text"\""}$ \\
& $|$ & \dots \\
& $|$ & $C_k \ @{text"\""}\tau_{k,1}@{text"\""} \dots @{text"\""}\tau_{k,n_k}@{text"\""}$
\end{tabular}
\end{quote}
It introduces the constructors \
-$C_i :: \tau_{i,1}\Rightarrow \cdots \Rightarrow \tau_{i,n_i} \Rightarrow$~@{text "('a\<^isub>1,\<dots>,'a\<^isub>n)t"} \ and expresses that any value of this type is built from these constructors in a unique manner. Uniqueness is implied by the following
+$C_i :: \tau_{i,1}\Rightarrow \cdots \Rightarrow \tau_{i,n_i} \Rightarrow$~@{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"} \ and expresses that any value of this type is built from these constructors in a unique manner. Uniqueness is implied by the following
properties of the constructors:
\begin{itemize}
\item \emph{Distinctness:} $C_i\ \ldots \neq C_j\ \dots$ \quad if $i \neq j$
@@ -40,9 +40,9 @@
\end{itemize}
The fact that any value of the datatype is built from the constructors implies
the structural induction rule: to show
-$P~x$ for all $x$ of type @{text "('a\<^isub>1,\<dots>,'a\<^isub>n)t"},
+$P~x$ for all $x$ of type @{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"},
one needs to show $P(C_i\ x_1 \dots x_{n_i})$ (for each $i$) assuming
-$P(x_j)$ for all $j$ where $\tau_{i,j} =$~@{text "('a\<^isub>1,\<dots>,'a\<^isub>n)t"}.
+$P(x_j)$ for all $j$ where $\tau_{i,j} =$~@{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"}.
Distinctness and injectivity are applied automatically by @{text auto}
and other proof methods. Induction must be applied explicitly.
@@ -91,11 +91,11 @@
"lookup ((a,b) # ps) x = (if a = x then Some b else lookup ps x)"
text{*
-Note that @{text"\<tau>\<^isub>1 * \<tau>\<^isub>2"} is the type of pairs, also written @{text"\<tau>\<^isub>1 \<times> \<tau>\<^isub>2"}.
+Note that @{text"\<tau>\<^sub>1 * \<tau>\<^sub>2"} is the type of pairs, also written @{text"\<tau>\<^sub>1 \<times> \<tau>\<^sub>2"}.
Pairs can be taken apart either by pattern matching (as above) or with the
projection functions @{const fst} and @{const snd}: @{thm fst_conv} and @{thm snd_conv}. Tuples are simulated by pairs nested to the right: @{term"(a,b,c)"}
-abbreviates @{text"(a, (b, c))"} and @{text "\<tau>\<^isub>1 \<times> \<tau>\<^isub>2 \<times> \<tau>\<^isub>3"} abbreviates
-@{text "\<tau>\<^isub>1 \<times> (\<tau>\<^isub>2 \<times> \<tau>\<^isub>3)"}.
+abbreviates @{text"(a, (b, c))"} and @{text "\<tau>\<^sub>1 \<times> \<tau>\<^sub>2 \<times> \<tau>\<^sub>3"} abbreviates
+@{text "\<tau>\<^sub>1 \<times> (\<tau>\<^sub>2 \<times> \<tau>\<^sub>3)"}.
\subsection{Definitions}
@@ -106,7 +106,7 @@
"sq n = n * n"
text{* Such definitions do not allow pattern matching but only
-@{text"f x\<^isub>1 \<dots> x\<^isub>n = t"}, where @{text f} does not occur in @{text t}.
+@{text"f x\<^sub>1 \<dots> x\<^sub>n = t"}, where @{text f} does not occur in @{text t}.
\subsection{Abbreviations}
@@ -185,18 +185,18 @@
because the induction follows the (terminating!) computation.
For every defining equation
\begin{quote}
-@{text "f(e) = \<dots> f(r\<^isub>1) \<dots> f(r\<^isub>k) \<dots>"}
+@{text "f(e) = \<dots> f(r\<^sub>1) \<dots> f(r\<^sub>k) \<dots>"}
\end{quote}
-where @{text"f(r\<^isub>i)"}, @{text"i=1\<dots>k"}, are all the recursive calls,
+where @{text"f(r\<^sub>i)"}, @{text"i=1\<dots>k"}, are all the recursive calls,
the induction rule @{text"f.induct"} contains one premise of the form
\begin{quote}
-@{text"P(r\<^isub>1) \<Longrightarrow> \<dots> \<Longrightarrow> P(r\<^isub>k) \<Longrightarrow> P(e)"}
+@{text"P(r\<^sub>1) \<Longrightarrow> \<dots> \<Longrightarrow> P(r\<^sub>k) \<Longrightarrow> P(e)"}
\end{quote}
-If @{text "f :: \<tau>\<^isub>1 \<Rightarrow> \<dots> \<Rightarrow> \<tau>\<^isub>n \<Rightarrow> \<tau>"} then @{text"f.induct"} is applied like this:
+If @{text "f :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> \<tau>"} then @{text"f.induct"} is applied like this:
\begin{quote}
-\isacom{apply}@{text"(induction x\<^isub>1 \<dots> x\<^isub>n rule: f.induct)"}
+\isacom{apply}@{text"(induction x\<^sub>1 \<dots> x\<^sub>n rule: f.induct)"}
\end{quote}
-where typically there is a call @{text"f x\<^isub>1 \<dots> x\<^isub>n"} in the goal.
+where typically there is a call @{text"f x\<^sub>1 \<dots> x\<^sub>n"} in the goal.
But note that the induction rule does not mention @{text f} at all,
except in its name, and is applicable independently of @{text f}.
@@ -300,7 +300,7 @@
\emph{Generalize induction by generalizing all free
variables\\ {\em(except the induction variable itself)}.}
\end{quote}
-Generalization is best performed with @{text"arbitrary: y\<^isub>1 \<dots> y\<^isub>k"}.
+Generalization is best performed with @{text"arbitrary: y\<^sub>1 \<dots> y\<^sub>k"}.
This heuristic prevents trivial failures like the one above.
However, it should not be applied blindly.
It is not always required, and the additional quantifiers can complicate
@@ -410,16 +410,16 @@
In such cases the more predictable @{text simp} method should be used.
Given a goal
\begin{quote}
-@{text"1. \<lbrakk> P\<^isub>1; \<dots>; P\<^isub>m \<rbrakk> \<Longrightarrow> C"}
+@{text"1. \<lbrakk> P\<^sub>1; \<dots>; P\<^sub>m \<rbrakk> \<Longrightarrow> C"}
\end{quote}
the command
\begin{quote}
-\isacom{apply}@{text"(simp add: th\<^isub>1 \<dots> th\<^isub>n)"}
+\isacom{apply}@{text"(simp add: th\<^sub>1 \<dots> th\<^sub>n)"}
\end{quote}
-simplifies the assumptions @{text "P\<^isub>i"} and the conclusion @{text C} using
+simplifies the assumptions @{text "P\<^sub>i"} and the conclusion @{text C} using
\begin{itemize}
\item all simplification rules, including the ones coming from \isacom{datatype} and \isacom{fun},
-\item the additional lemmas @{text"th\<^isub>1 \<dots> th\<^isub>n"}, and
+\item the additional lemmas @{text"th\<^sub>1 \<dots> th\<^sub>n"}, and
\item the assumptions.
\end{itemize}
In addition to or instead of @{text add} there is also @{text del} for removing
--- a/src/Doc/Tutorial/Documents/Documents.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Tutorial/Documents/Documents.thy Tue Aug 13 16:25:47 2013 +0200
@@ -126,10 +126,10 @@
in the trailing part of an identifier. This means that the input
\medskip
- {\small\noindent \verb,\,\verb,<forall>\,\verb,<alpha>\<^isub>1.,~\verb,\,\verb,<alpha>\<^isub>1 = \,\verb,<Pi>\<^isub>\<A>,}
+ {\small\noindent \verb,\,\verb,<forall>\,\verb,<alpha>\<^sub>1.,~\verb,\,\verb,<alpha>\<^sub>1 = \,\verb,<Pi>\<^sub>\<A>,}
\medskip
- \noindent is recognized as the term @{term "\<forall>\<alpha>\<^isub>1. \<alpha>\<^isub>1 = \<Pi>\<^isub>\<A>"}
+ \noindent is recognized as the term @{term "\<forall>\<alpha>\<^sub>1. \<alpha>\<^sub>1 = \<Pi>\<^sub>\<A>"}
by Isabelle.
Replacing our previous definition of @{text xor} by the
--- a/src/Doc/Tutorial/Inductive/Mutual.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Tutorial/Inductive/Mutual.thy Tue Aug 13 16:25:47 2013 +0200
@@ -69,11 +69,11 @@
you write \commdx{inductive} instead of \isacommand{inductive\_set} and
@{prop"evn n"} instead of @{prop"n : even"}.
When defining an n-ary relation as a predicate, it is recommended to curry
-the predicate: its type should be \mbox{@{text"\<tau>\<^isub>1 \<Rightarrow> \<dots> \<Rightarrow> \<tau>\<^isub>n \<Rightarrow> bool"}}
+the predicate: its type should be \mbox{@{text"\<tau>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> bool"}}
rather than
-@{text"\<tau>\<^isub>1 \<times> \<dots> \<times> \<tau>\<^isub>n \<Rightarrow> bool"}. The curried version facilitates inductions.
+@{text"\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<Rightarrow> bool"}. The curried version facilitates inductions.
-When should you choose sets and when predicates? If you intend to combine your notion with set theoretic notation, define it as an inductive set. If not, define it as an inductive predicate, thus avoiding the @{text"\<in>"} notation. But note that predicates of more than one argument cannot be combined with the usual set theoretic operators: @{term"P \<union> Q"} is not well-typed if @{text"P, Q :: \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2 \<Rightarrow> bool"}, you have to write @{term"%x y. P x y & Q x y"} instead.
+When should you choose sets and when predicates? If you intend to combine your notion with set theoretic notation, define it as an inductive set. If not, define it as an inductive predicate, thus avoiding the @{text"\<in>"} notation. But note that predicates of more than one argument cannot be combined with the usual set theoretic operators: @{term"P \<union> Q"} is not well-typed if @{text"P, Q :: \<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2 \<Rightarrow> bool"}, you have to write @{term"%x y. P x y & Q x y"} instead.
\index{inductive predicates|)}
*}
--- a/src/Doc/Tutorial/Types/Axioms.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/Doc/Tutorial/Types/Axioms.thy Tue Aug 13 16:25:47 2013 +0200
@@ -63,9 +63,9 @@
begin
instance proof
- fix p\<^isub>1 p\<^isub>2 p\<^isub>3 :: "'a\<Colon>semigroup \<times> 'b\<Colon>semigroup"
- show "p\<^isub>1 \<oplus> p\<^isub>2 \<oplus> p\<^isub>3 = p\<^isub>1 \<oplus> (p\<^isub>2 \<oplus> p\<^isub>3)"
- by (cases p\<^isub>1, cases p\<^isub>2, cases p\<^isub>3) (simp add: assoc)
+ fix p\<^sub>1 p\<^sub>2 p\<^sub>3 :: "'a\<Colon>semigroup \<times> 'b\<Colon>semigroup"
+ show "p\<^sub>1 \<oplus> p\<^sub>2 \<oplus> p\<^sub>3 = p\<^sub>1 \<oplus> (p\<^sub>2 \<oplus> p\<^sub>3)"
+ by (cases p\<^sub>1, cases p\<^sub>2, cases p\<^sub>3) (simp add: assoc)
txt {* \noindent Associativity of product semigroups is established
using the hypothetical associativity @{fact assoc} of the type
--- a/src/HOL/Complete_Lattices.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Complete_Lattices.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1033,8 +1033,8 @@
text {*
Note the difference between ordinary xsymbol syntax of indexed
- unions and intersections (e.g.\ @{text"\<Union>a\<^isub>1\<in>A\<^isub>1. B"})
- and their \LaTeX\ rendition: @{term"\<Union>a\<^isub>1\<in>A\<^isub>1. B"}. The
+ unions and intersections (e.g.\ @{text"\<Union>a\<^sub>1\<in>A\<^sub>1. B"})
+ and their \LaTeX\ rendition: @{term"\<Union>a\<^sub>1\<in>A\<^sub>1. B"}. The
former does not make the index expression a subscript of the
union/intersection symbol because this leads to problems with nested
subscripts in Proof General.
--- a/src/HOL/Complex.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Complex.thy Tue Aug 13 16:25:47 2013 +0200
@@ -103,7 +103,7 @@
definition complex_inverse_def:
"inverse x =
- Complex (Re x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)) (- Im x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>))"
+ Complex (Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)) (- Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2))"
definition complex_divide_def:
"x / (y\<Colon>complex) = x * inverse y"
@@ -128,15 +128,15 @@
by (simp add: complex_mult_def)
lemma complex_inverse [simp]:
- "inverse (Complex a b) = Complex (a / (a\<twosuperior> + b\<twosuperior>)) (- b / (a\<twosuperior> + b\<twosuperior>))"
+ "inverse (Complex a b) = Complex (a / (a\<^sup>2 + b\<^sup>2)) (- b / (a\<^sup>2 + b\<^sup>2))"
by (simp add: complex_inverse_def)
lemma complex_Re_inverse:
- "Re (inverse x) = Re x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)"
+ "Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)"
by (simp add: complex_inverse_def)
lemma complex_Im_inverse:
- "Im (inverse x) = - Im x / ((Re x)\<twosuperior> + (Im x)\<twosuperior>)"
+ "Im (inverse x) = - Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)"
by (simp add: complex_inverse_def)
instance
@@ -267,7 +267,7 @@
begin
definition complex_norm_def:
- "norm z = sqrt ((Re z)\<twosuperior> + (Im z)\<twosuperior>)"
+ "norm z = sqrt ((Re z)\<^sup>2 + (Im z)\<^sup>2)"
abbreviation cmod :: "complex \<Rightarrow> real"
where "cmod \<equiv> norm"
@@ -283,7 +283,7 @@
lemmas cmod_def = complex_norm_def
-lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma complex_norm [simp]: "cmod (Complex x y) = sqrt (x\<^sup>2 + y\<^sup>2)"
by (simp add: complex_norm_def)
instance proof
@@ -439,7 +439,7 @@
lemma i_squared [simp]: "ii * ii = -1"
by (simp add: i_def)
-lemma power2_i [simp]: "ii\<twosuperior> = -1"
+lemma power2_i [simp]: "ii\<^sup>2 = -1"
by (simp add: power2_eq_square)
lemma inverse_i [simp]: "inverse ii = - ii"
@@ -529,10 +529,10 @@
lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * ii"
by (simp add: complex_eq_iff)
-lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<twosuperior> + (Im z)\<twosuperior>)"
+lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<^sup>2 + (Im z)\<^sup>2)"
by (simp add: complex_eq_iff power2_eq_square)
-lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<twosuperior>"
+lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<^sup>2"
by (simp add: norm_mult power2_eq_square)
lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))"
--- a/src/HOL/Decision_Procs/Cooper.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1110,7 +1110,7 @@
"plusinf (Ge (CN 0 c e)) = T"
"plusinf p = p"
-consts \<delta> :: "fm \<Rightarrow> int" -- {* Compute @{text "lcm {d| N\<^isup>? Dvd c*x+t \<in> p}"} *}
+consts \<delta> :: "fm \<Rightarrow> int" -- {* Compute @{text "lcm {d| N\<^sup>? Dvd c*x+t \<in> p}"} *}
recdef \<delta> "measure size"
"\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
"\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
--- a/src/HOL/Decision_Procs/ex/Approximation_Ex.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Decision_Procs/ex/Approximation_Ex.thy Tue Aug 13 16:25:47 2013 +0200
@@ -17,8 +17,8 @@
variables can be used, but each one need to be bounded by an upper and lower
bound.
-To specify the bounds either @{term "l\<^isub>1 \<le> x \<and> x \<le> u\<^isub>1"},
-@{term "x \<in> { l\<^isub>1 .. u\<^isub>1 }"} or @{term "x = bnd"} can be used. Where the
+To specify the bounds either @{term "l\<^sub>1 \<le> x \<and> x \<le> u\<^sub>1"},
+@{term "x \<in> { l\<^sub>1 .. u\<^sub>1 }"} or @{term "x = bnd"} can be used. Where the
bound specification are again arithmetic formulas containing variables. They can
be connected using either meta level or HOL equivalence.
--- a/src/HOL/Enum.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Enum.thy Tue Aug 13 16:25:47 2013 +0200
@@ -438,25 +438,25 @@
text {* We define small finite types for the use in Quickcheck *}
-datatype finite_1 = a\<^isub>1
+datatype finite_1 = a\<^sub>1
-notation (output) a\<^isub>1 ("a\<^isub>1")
+notation (output) a\<^sub>1 ("a\<^sub>1")
lemma UNIV_finite_1:
- "UNIV = {a\<^isub>1}"
+ "UNIV = {a\<^sub>1}"
by (auto intro: finite_1.exhaust)
instantiation finite_1 :: enum
begin
definition
- "enum = [a\<^isub>1]"
+ "enum = [a\<^sub>1]"
definition
- "enum_all P = P a\<^isub>1"
+ "enum_all P = P a\<^sub>1"
definition
- "enum_ex P = P a\<^isub>1"
+ "enum_ex P = P a\<^sub>1"
instance proof
qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all)
@@ -482,28 +482,28 @@
end
-hide_const (open) a\<^isub>1
+hide_const (open) a\<^sub>1
-datatype finite_2 = a\<^isub>1 | a\<^isub>2
+datatype finite_2 = a\<^sub>1 | a\<^sub>2
-notation (output) a\<^isub>1 ("a\<^isub>1")
-notation (output) a\<^isub>2 ("a\<^isub>2")
+notation (output) a\<^sub>1 ("a\<^sub>1")
+notation (output) a\<^sub>2 ("a\<^sub>2")
lemma UNIV_finite_2:
- "UNIV = {a\<^isub>1, a\<^isub>2}"
+ "UNIV = {a\<^sub>1, a\<^sub>2}"
by (auto intro: finite_2.exhaust)
instantiation finite_2 :: enum
begin
definition
- "enum = [a\<^isub>1, a\<^isub>2]"
+ "enum = [a\<^sub>1, a\<^sub>2]"
definition
- "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2"
+ "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2"
definition
- "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2"
+ "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2"
instance proof
qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all)
@@ -515,7 +515,7 @@
definition less_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
where
- "x < y \<longleftrightarrow> x = a\<^isub>1 \<and> y = a\<^isub>2"
+ "x < y \<longleftrightarrow> x = a\<^sub>1 \<and> y = a\<^sub>2"
definition less_eq_finite_2 :: "finite_2 \<Rightarrow> finite_2 \<Rightarrow> bool"
where
@@ -529,29 +529,29 @@
end
-hide_const (open) a\<^isub>1 a\<^isub>2
+hide_const (open) a\<^sub>1 a\<^sub>2
-datatype finite_3 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3
+datatype finite_3 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3
-notation (output) a\<^isub>1 ("a\<^isub>1")
-notation (output) a\<^isub>2 ("a\<^isub>2")
-notation (output) a\<^isub>3 ("a\<^isub>3")
+notation (output) a\<^sub>1 ("a\<^sub>1")
+notation (output) a\<^sub>2 ("a\<^sub>2")
+notation (output) a\<^sub>3 ("a\<^sub>3")
lemma UNIV_finite_3:
- "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3}"
+ "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3}"
by (auto intro: finite_3.exhaust)
instantiation finite_3 :: enum
begin
definition
- "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3]"
+ "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3]"
definition
- "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3"
+ "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3"
definition
- "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3"
+ "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3"
instance proof
qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all)
@@ -563,7 +563,7 @@
definition less_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
where
- "x < y = (case x of a\<^isub>1 \<Rightarrow> y \<noteq> a\<^isub>1 | a\<^isub>2 \<Rightarrow> y = a\<^isub>3 | a\<^isub>3 \<Rightarrow> False)"
+ "x < y = (case x of a\<^sub>1 \<Rightarrow> y \<noteq> a\<^sub>1 | a\<^sub>2 \<Rightarrow> y = a\<^sub>3 | a\<^sub>3 \<Rightarrow> False)"
definition less_eq_finite_3 :: "finite_3 \<Rightarrow> finite_3 \<Rightarrow> bool"
where
@@ -574,69 +574,69 @@
end
-hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3
+hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3
-datatype finite_4 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4
+datatype finite_4 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4
-notation (output) a\<^isub>1 ("a\<^isub>1")
-notation (output) a\<^isub>2 ("a\<^isub>2")
-notation (output) a\<^isub>3 ("a\<^isub>3")
-notation (output) a\<^isub>4 ("a\<^isub>4")
+notation (output) a\<^sub>1 ("a\<^sub>1")
+notation (output) a\<^sub>2 ("a\<^sub>2")
+notation (output) a\<^sub>3 ("a\<^sub>3")
+notation (output) a\<^sub>4 ("a\<^sub>4")
lemma UNIV_finite_4:
- "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4}"
+ "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4}"
by (auto intro: finite_4.exhaust)
instantiation finite_4 :: enum
begin
definition
- "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4]"
+ "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4]"
definition
- "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4"
+ "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4"
definition
- "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4"
+ "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4"
instance proof
qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all)
end
-hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4
+hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4
-datatype finite_5 = a\<^isub>1 | a\<^isub>2 | a\<^isub>3 | a\<^isub>4 | a\<^isub>5
+datatype finite_5 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 | a\<^sub>5
-notation (output) a\<^isub>1 ("a\<^isub>1")
-notation (output) a\<^isub>2 ("a\<^isub>2")
-notation (output) a\<^isub>3 ("a\<^isub>3")
-notation (output) a\<^isub>4 ("a\<^isub>4")
-notation (output) a\<^isub>5 ("a\<^isub>5")
+notation (output) a\<^sub>1 ("a\<^sub>1")
+notation (output) a\<^sub>2 ("a\<^sub>2")
+notation (output) a\<^sub>3 ("a\<^sub>3")
+notation (output) a\<^sub>4 ("a\<^sub>4")
+notation (output) a\<^sub>5 ("a\<^sub>5")
lemma UNIV_finite_5:
- "UNIV = {a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5}"
+ "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5}"
by (auto intro: finite_5.exhaust)
instantiation finite_5 :: enum
begin
definition
- "enum = [a\<^isub>1, a\<^isub>2, a\<^isub>3, a\<^isub>4, a\<^isub>5]"
+ "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5]"
definition
- "enum_all P \<longleftrightarrow> P a\<^isub>1 \<and> P a\<^isub>2 \<and> P a\<^isub>3 \<and> P a\<^isub>4 \<and> P a\<^isub>5"
+ "enum_all P \<longleftrightarrow> P a\<^sub>1 \<and> P a\<^sub>2 \<and> P a\<^sub>3 \<and> P a\<^sub>4 \<and> P a\<^sub>5"
definition
- "enum_ex P \<longleftrightarrow> P a\<^isub>1 \<or> P a\<^isub>2 \<or> P a\<^isub>3 \<or> P a\<^isub>4 \<or> P a\<^isub>5"
+ "enum_ex P \<longleftrightarrow> P a\<^sub>1 \<or> P a\<^sub>2 \<or> P a\<^sub>3 \<or> P a\<^sub>4 \<or> P a\<^sub>5"
instance proof
qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all)
end
-hide_const (open) a\<^isub>1 a\<^isub>2 a\<^isub>3 a\<^isub>4 a\<^isub>5
+hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5
subsection {* Closing up *}
--- a/src/HOL/Finite_Set.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Finite_Set.thy Tue Aug 13 16:25:47 2013 +0200
@@ -566,7 +566,7 @@
subsection {* A basic fold functional for finite sets *}
text {* The intended behaviour is
-@{text "fold f z {x\<^isub>1, ..., x\<^isub>n} = f x\<^isub>1 (\<dots> (f x\<^isub>n z)\<dots>)"}
+@{text "fold f z {x\<^sub>1, ..., x\<^sub>n} = f x\<^sub>1 (\<dots> (f x\<^sub>n z)\<dots>)"}
if @{text f} is ``left-commutative'':
*}
--- a/src/HOL/IMP/ACom.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/ACom.thy Tue Aug 13 16:25:47 2013 +0200
@@ -19,9 +19,9 @@
fun strip :: "'a acom \<Rightarrow> com" where
"strip (SKIP {P}) = com.SKIP" |
"strip (x ::= e {P}) = x ::= e" |
-"strip (C\<^isub>1;;C\<^isub>2) = strip C\<^isub>1;; strip C\<^isub>2" |
-"strip (IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {P}) =
- IF b THEN strip C\<^isub>1 ELSE strip C\<^isub>2" |
+"strip (C\<^sub>1;;C\<^sub>2) = strip C\<^sub>1;; strip C\<^sub>2" |
+"strip (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {P}) =
+ IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2" |
"strip ({I} WHILE b DO {P} C {Q}) = WHILE b DO strip C"
text_raw{*}%endsnip*}
@@ -29,8 +29,8 @@
fun asize :: "com \<Rightarrow> nat" where
"asize com.SKIP = 1" |
"asize (x ::= e) = 1" |
-"asize (C\<^isub>1;;C\<^isub>2) = asize C\<^isub>1 + asize C\<^isub>2" |
-"asize (IF b THEN C\<^isub>1 ELSE C\<^isub>2) = asize C\<^isub>1 + asize C\<^isub>2 + 3" |
+"asize (C\<^sub>1;;C\<^sub>2) = asize C\<^sub>1 + asize C\<^sub>2" |
+"asize (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = asize C\<^sub>1 + asize C\<^sub>2 + 3" |
"asize (WHILE b DO C) = asize C + 3"
text_raw{*}%endsnip*}
@@ -41,11 +41,11 @@
fun annotate :: "(nat \<Rightarrow> 'a) \<Rightarrow> com \<Rightarrow> 'a acom" where
"annotate f com.SKIP = SKIP {f 0}" |
"annotate f (x ::= e) = x ::= e {f 0}" |
-"annotate f (c\<^isub>1;;c\<^isub>2) = annotate f c\<^isub>1;; annotate (shift f (asize c\<^isub>1)) c\<^isub>2" |
-"annotate f (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
- IF b THEN {f 0} annotate (shift f 1) c\<^isub>1
- ELSE {f(asize c\<^isub>1 + 1)} annotate (shift f (asize c\<^isub>1 + 2)) c\<^isub>2
- {f(asize c\<^isub>1 + asize c\<^isub>2 + 2)}" |
+"annotate f (c\<^sub>1;;c\<^sub>2) = annotate f c\<^sub>1;; annotate (shift f (asize c\<^sub>1)) c\<^sub>2" |
+"annotate f (IF b THEN c\<^sub>1 ELSE c\<^sub>2) =
+ IF b THEN {f 0} annotate (shift f 1) c\<^sub>1
+ ELSE {f(asize c\<^sub>1 + 1)} annotate (shift f (asize c\<^sub>1 + 2)) c\<^sub>2
+ {f(asize c\<^sub>1 + asize c\<^sub>2 + 2)}" |
"annotate f (WHILE b DO c) =
{f 0} WHILE b DO {f 1} annotate (shift f 2) c {f(asize c + 2)}"
text_raw{*}%endsnip*}
@@ -54,9 +54,9 @@
fun annos :: "'a acom \<Rightarrow> 'a list" where
"annos (SKIP {P}) = [P]" |
"annos (x ::= e {P}) = [P]" |
-"annos (C\<^isub>1;;C\<^isub>2) = annos C\<^isub>1 @ annos C\<^isub>2" |
-"annos (IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) =
- P\<^isub>1 # annos C\<^isub>1 @ P\<^isub>2 # annos C\<^isub>2 @ [Q]" |
+"annos (C\<^sub>1;;C\<^sub>2) = annos C\<^sub>1 @ annos C\<^sub>2" |
+"annos (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {Q}) =
+ P\<^sub>1 # annos C\<^sub>1 @ P\<^sub>2 # annos C\<^sub>2 @ [Q]" |
"annos ({I} WHILE b DO {P} C {Q}) = I # P # annos C @ [Q]"
text_raw{*}%endsnip*}
@@ -70,9 +70,9 @@
fun map_acom :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a acom \<Rightarrow> 'b acom" where
"map_acom f (SKIP {P}) = SKIP {f P}" |
"map_acom f (x ::= e {P}) = x ::= e {f P}" |
-"map_acom f (C\<^isub>1;;C\<^isub>2) = map_acom f C\<^isub>1;; map_acom f C\<^isub>2" |
-"map_acom f (IF b THEN {P\<^isub>1} C\<^isub>1 ELSE {P\<^isub>2} C\<^isub>2 {Q}) =
- IF b THEN {f P\<^isub>1} map_acom f C\<^isub>1 ELSE {f P\<^isub>2} map_acom f C\<^isub>2
+"map_acom f (C\<^sub>1;;C\<^sub>2) = map_acom f C\<^sub>1;; map_acom f C\<^sub>2" |
+"map_acom f (IF b THEN {P\<^sub>1} C\<^sub>1 ELSE {P\<^sub>2} C\<^sub>2 {Q}) =
+ IF b THEN {f P\<^sub>1} map_acom f C\<^sub>1 ELSE {f P\<^sub>2} map_acom f C\<^sub>2
{f Q}" |
"map_acom f ({I} WHILE b DO {P} C {Q}) =
{f I} WHILE b DO {f P} map_acom f C {f Q}"
--- a/src/HOL/IMP/AExp.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/AExp.thy Tue Aug 13 16:25:47 2013 +0200
@@ -16,7 +16,7 @@
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val" where
"aval (N n) s = n" |
"aval (V x) s = s x" |
-"aval (Plus a\<^isub>1 a\<^isub>2) s = aval a\<^isub>1 s + aval a\<^isub>2 s"
+"aval (Plus a\<^sub>1 a\<^sub>2) s = aval a\<^sub>1 s + aval a\<^sub>2 s"
text_raw{*}%endsnip*}
@@ -48,7 +48,7 @@
value "aval (Plus (V ''x'') (N 5)) <''y'' := 7>"
text{* Note that this @{text"<\<dots>>"} syntax works for any function space
-@{text"\<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2"} where @{text "\<tau>\<^isub>2"} has a @{text 0}. *}
+@{text"\<tau>\<^sub>1 \<Rightarrow> \<tau>\<^sub>2"} where @{text "\<tau>\<^sub>2"} has a @{text 0}. *}
subsection "Constant Folding"
@@ -59,10 +59,10 @@
fun asimp_const :: "aexp \<Rightarrow> aexp" where
"asimp_const (N n) = N n" |
"asimp_const (V x) = V x" |
-"asimp_const (Plus a\<^isub>1 a\<^isub>2) =
- (case (asimp_const a\<^isub>1, asimp_const a\<^isub>2) of
- (N n\<^isub>1, N n\<^isub>2) \<Rightarrow> N(n\<^isub>1+n\<^isub>2) |
- (b\<^isub>1,b\<^isub>2) \<Rightarrow> Plus b\<^isub>1 b\<^isub>2)"
+"asimp_const (Plus a\<^sub>1 a\<^sub>2) =
+ (case (asimp_const a\<^sub>1, asimp_const a\<^sub>2) of
+ (N n\<^sub>1, N n\<^sub>2) \<Rightarrow> N(n\<^sub>1+n\<^sub>2) |
+ (b\<^sub>1,b\<^sub>2) \<Rightarrow> Plus b\<^sub>1 b\<^sub>2)"
text_raw{*}%endsnip*}
theorem aval_asimp_const:
@@ -76,10 +76,10 @@
text_raw{*\snip{AExpplusdef}{0}{2}{% *}
fun plus :: "aexp \<Rightarrow> aexp \<Rightarrow> aexp" where
-"plus (N i\<^isub>1) (N i\<^isub>2) = N(i\<^isub>1+i\<^isub>2)" |
+"plus (N i\<^sub>1) (N i\<^sub>2) = N(i\<^sub>1+i\<^sub>2)" |
"plus (N i) a = (if i=0 then a else Plus (N i) a)" |
"plus a (N i) = (if i=0 then a else Plus a (N i))" |
-"plus a\<^isub>1 a\<^isub>2 = Plus a\<^isub>1 a\<^isub>2"
+"plus a\<^sub>1 a\<^sub>2 = Plus a\<^sub>1 a\<^sub>2"
text_raw{*}%endsnip*}
lemma aval_plus[simp]:
@@ -92,7 +92,7 @@
fun asimp :: "aexp \<Rightarrow> aexp" where
"asimp (N n) = N n" |
"asimp (V x) = V x" |
-"asimp (Plus a\<^isub>1 a\<^isub>2) = plus (asimp a\<^isub>1) (asimp a\<^isub>2)"
+"asimp (Plus a\<^sub>1 a\<^sub>2) = plus (asimp a\<^sub>1) (asimp a\<^sub>2)"
text_raw{*}%endsnip*}
text{* Note that in @{const asimp_const} the optimized constructor was
--- a/src/HOL/IMP/ASM.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/ASM.thy Tue Aug 13 16:25:47 2013 +0200
@@ -47,7 +47,7 @@
fun comp :: "aexp \<Rightarrow> instr list" where
"comp (N n) = [LOADI n]" |
"comp (V x) = [LOAD x]" |
-"comp (Plus e\<^isub>1 e\<^isub>2) = comp e\<^isub>1 @ comp e\<^isub>2 @ [ADD]"
+"comp (Plus e\<^sub>1 e\<^sub>2) = comp e\<^sub>1 @ comp e\<^sub>2 @ [ADD]"
text_raw{*}%endsnip*}
value "comp (Plus (Plus (V ''x'') (N 1)) (V ''z''))"
--- a/src/HOL/IMP/Abs_Int0.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int0.thy Tue Aug 13 16:25:47 2013 +0200
@@ -183,63 +183,63 @@
"AI c = pfp (step' \<top>) (bot c)"
-abbreviation \<gamma>\<^isub>s :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>s == \<gamma>_fun \<gamma>"
+abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^sub>s == \<gamma>_fun \<gamma>"
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>s"
+abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s"
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
-lemma gamma_s_Top[simp]: "\<gamma>\<^isub>s \<top> = UNIV"
+lemma gamma_s_Top[simp]: "\<gamma>\<^sub>s \<top> = UNIV"
by(simp add: top_fun_def \<gamma>_fun_def)
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o \<top> = UNIV"
+lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV"
by (simp add: top_option_def)
-lemma mono_gamma_s: "f1 \<le> f2 \<Longrightarrow> \<gamma>\<^isub>s f1 \<subseteq> \<gamma>\<^isub>s f2"
+lemma mono_gamma_s: "f1 \<le> f2 \<Longrightarrow> \<gamma>\<^sub>s f1 \<subseteq> \<gamma>\<^sub>s f2"
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
lemma mono_gamma_o:
- "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^isub>o S1 \<subseteq> \<gamma>\<^isub>o S2"
+ "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2"
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s)
-lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^isub>c C1 \<le> \<gamma>\<^isub>c C2"
+lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2"
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2])
text{* Correctness: *}
-lemma aval'_correct: "s : \<gamma>\<^isub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+lemma aval'_correct: "s : \<gamma>\<^sub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
-lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(S(x := a))"
+lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^sub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>s(S(x := a))"
by(simp add: \<gamma>_fun_def)
lemma gamma_Step_subcomm:
- assumes "!!x e S. f1 x e (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (f2 x e S)" "!!b S. g1 b (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (g2 b S)"
- shows "Step f1 g1 (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (Step f2 g2 S C)"
+ assumes "!!x e S. f1 x e (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (f2 x e S)" "!!b S. g1 b (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (g2 b S)"
+ shows "Step f1 g1 (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (Step f2 g2 S C)"
proof(induction C arbitrary: S)
qed (auto simp: mono_gamma_o assms)
-lemma step_step': "step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)"
+lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)"
unfolding step_def step'_def
by(rule gamma_Step_subcomm)
(auto simp: aval'_correct in_gamma_update asem_def split: option.splits)
-lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
+lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
proof(simp add: CS_def AI_def)
assume 1: "pfp (step' \<top>) (bot c) = Some C"
have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])
- have 2: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C" --"transfer the pfp'"
+ have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C" --"transfer the pfp'"
proof(rule order_trans)
- show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" by(rule step_step')
- show "... \<le> \<gamma>\<^isub>c C" by (metis mono_gamma_c[OF pfp'])
+ show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)" by(rule step_step')
+ show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp'])
qed
- have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
- have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C"
- by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 2])
- thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
+ have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
+ have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
+ by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 2])
+ thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
qed
end
@@ -314,20 +314,20 @@
assumes h: "m x \<le> h"
begin
-definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^isub>s") where
+definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
"m_s S X = (\<Sum> x \<in> X. m(S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"
by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h])
-fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^isub>o") where
+fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
"m_o (Some S) X = m_s S X" |
"m_o None X = h * card X + 1"
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (h*card X + 1)"
by(cases opt)(auto simp add: m_s_h le_SucI dest: m_s_h)
-definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^isub>c") where
+definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^sub>c") where
"m_c C = listsum (map (\<lambda>a. m_o a (vars C)) (annos C))"
text{* Upper complexity bound: *}
@@ -356,14 +356,14 @@
the finitely many variables in the program change. That the others do not change
follows because they remain @{term \<top>}. *}
-fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>s") where
+fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
"top_on_st S X = (\<forall>x\<in>X. S x = \<top>)"
-fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>o") where
+fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
"top_on_opt (Some S) X = top_on_st S X" |
"top_on_opt None X = True"
-definition top_on_acom :: "'av st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>c") where
+definition top_on_acom :: "'av st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"
lemma top_on_top: "top_on_opt \<top> X"
--- a/src/HOL/IMP/Abs_Int1.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int1.thy Tue Aug 13 16:25:47 2013 +0200
@@ -16,17 +16,17 @@
"aval' (V x) S = fun S x" |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
-lemma aval'_correct: "s : \<gamma>\<^isub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+lemma aval'_correct: "s : \<gamma>\<^sub>s S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def)
lemma gamma_Step_subcomm: fixes C1 C2 :: "'a::semilattice_sup acom"
- assumes "!!x e S. f1 x e (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (f2 x e S)"
- "!!b S. g1 b (\<gamma>\<^isub>o S) \<subseteq> \<gamma>\<^isub>o (g2 b S)"
- shows "Step f1 g1 (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (Step f2 g2 S C)"
+ assumes "!!x e S. f1 x e (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (f2 x e S)"
+ "!!b S. g1 b (\<gamma>\<^sub>o S) \<subseteq> \<gamma>\<^sub>o (g2 b S)"
+ shows "Step f1 g1 (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (Step f2 g2 S C)"
proof(induction C arbitrary: S)
qed (auto simp: assms intro!: mono_gamma_o sup_ge1 sup_ge2)
-lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^isub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>s(update S x a)"
+lemma in_gamma_update: "\<lbrakk> s : \<gamma>\<^sub>s S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>s(update S x a)"
by(simp add: \<gamma>_st_def)
end
@@ -50,24 +50,24 @@
text{* Correctness: *}
-lemma step_step': "step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)"
+lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)"
unfolding step_def step'_def
by(rule gamma_Step_subcomm)
(auto simp: intro!: aval'_correct in_gamma_update split: option.splits)
-lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
+lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
proof(simp add: CS_def AI_def)
assume 1: "pfp (step' \<top>) (bot c) = Some C"
have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])
- have 2: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C" --"transfer the pfp'"
+ have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C" --"transfer the pfp'"
proof(rule order_trans)
- show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" by(rule step_step')
- show "... \<le> \<gamma>\<^isub>c C" by (metis mono_gamma_c[OF pfp'])
+ show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)" by(rule step_step')
+ show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp'])
qed
- have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
- have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C"
- by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 2])
- thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
+ have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
+ have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
+ by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 2])
+ thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
qed
end
@@ -105,19 +105,19 @@
assumes h: "m x \<le> h"
begin
-definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^isub>s") where
+definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
"m_s S X = (\<Sum> x \<in> X. m(fun S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"
by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h])
-definition m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^isub>o") where
+definition m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
"m_o opt X = (case opt of None \<Rightarrow> h * card X + 1 | Some S \<Rightarrow> m_s S X)"
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (h*card X + 1)"
by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h)
-definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^isub>c") where
+definition m_c :: "'av st option acom \<Rightarrow> nat" ("m\<^sub>c") where
"m_c C = listsum (map (\<lambda>a. m_o a (vars C)) (annos C))"
text{* Upper complexity bound: *}
@@ -134,14 +134,14 @@
end
-fun top_on_st :: "'a::order_top st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>s") where
+fun top_on_st :: "'a::order_top st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
"top_on_st S X = (\<forall>x\<in>X. fun S x = \<top>)"
-fun top_on_opt :: "'a::order_top st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>o") where
+fun top_on_opt :: "'a::order_top st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
"top_on_opt (Some S) X = top_on_st S X" |
"top_on_opt None X = True"
-definition top_on_acom :: "'a::order_top st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^isub>c") where
+definition top_on_acom :: "'a::order_top st option acom \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"
lemma top_on_top: "top_on_opt (\<top>::_ st option) X"
--- a/src/HOL/IMP/Abs_Int2.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int2.thy Tue Aug 13 16:25:47 2013 +0200
@@ -51,10 +51,10 @@
and inv_plus' :: "'av \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
and inv_less' :: "bool \<Rightarrow> 'av \<Rightarrow> 'av \<Rightarrow> 'av * 'av"
assumes test_num': "test_num' i a = (i : \<gamma> a)"
-and inv_plus': "inv_plus' a a1 a2 = (a\<^isub>1',a\<^isub>2') \<Longrightarrow>
- i1 : \<gamma> a1 \<Longrightarrow> i2 : \<gamma> a2 \<Longrightarrow> i1+i2 : \<gamma> a \<Longrightarrow> i1 : \<gamma> a\<^isub>1' \<and> i2 : \<gamma> a\<^isub>2'"
-and inv_less': "inv_less' (i1<i2) a1 a2 = (a\<^isub>1',a\<^isub>2') \<Longrightarrow>
- i1 : \<gamma> a1 \<Longrightarrow> i2 : \<gamma> a2 \<Longrightarrow> i1 : \<gamma> a\<^isub>1' \<and> i2 : \<gamma> a\<^isub>2'"
+and inv_plus': "inv_plus' a a1 a2 = (a\<^sub>1',a\<^sub>2') \<Longrightarrow>
+ i1 : \<gamma> a1 \<Longrightarrow> i2 : \<gamma> a2 \<Longrightarrow> i1+i2 : \<gamma> a \<Longrightarrow> i1 : \<gamma> a\<^sub>1' \<and> i2 : \<gamma> a\<^sub>2'"
+and inv_less': "inv_less' (i1<i2) a1 a2 = (a\<^sub>1',a\<^sub>2') \<Longrightarrow>
+ i1 : \<gamma> a1 \<Longrightarrow> i2 : \<gamma> a2 \<Longrightarrow> i1 : \<gamma> a\<^sub>1' \<and> i2 : \<gamma> a\<^sub>2'"
locale Abs_Int_inv = Val_inv where \<gamma> = \<gamma>
@@ -62,14 +62,14 @@
begin
lemma in_gamma_sup_UpI:
- "s : \<gamma>\<^isub>o S1 \<or> s : \<gamma>\<^isub>o S2 \<Longrightarrow> s : \<gamma>\<^isub>o(S1 \<squnion> S2)"
+ "s : \<gamma>\<^sub>o S1 \<or> s : \<gamma>\<^sub>o S2 \<Longrightarrow> s : \<gamma>\<^sub>o(S1 \<squnion> S2)"
by (metis (hide_lams, no_types) sup_ge1 sup_ge2 mono_gamma_o subsetD)
fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where
"aval'' e None = \<bottom>" |
"aval'' e (Some S) = aval' e S"
-lemma aval''_correct: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
+lemma aval''_correct: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
by(cases S)(auto simp add: aval'_correct split: option.splits)
subsubsection "Backward analysis"
@@ -103,12 +103,12 @@
(let (a1,a2) = inv_less' res (aval'' e1 S) (aval'' e2 S)
in inv_aval'' e1 a1 (inv_aval'' e2 a2 S))"
-lemma inv_aval''_correct: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^isub>o (inv_aval'' e a S)"
+lemma inv_aval''_correct: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^sub>o (inv_aval'' e a S)"
proof(induction e arbitrary: a S)
case N thus ?case by simp (metis test_num')
next
case (V x)
- obtain S' where "S = Some S'" and "s : \<gamma>\<^isub>s S'" using `s : \<gamma>\<^isub>o S`
+ obtain S' where "S = Some S'" and "s : \<gamma>\<^sub>s S'" using `s : \<gamma>\<^sub>o S`
by(auto simp: in_gamma_option_iff)
moreover hence "s x : \<gamma> (fun S' x)"
by(simp add: \<gamma>_st_def)
@@ -122,7 +122,7 @@
by (auto split: prod.split)
qed
-lemma inv_bval''_correct: "s : \<gamma>\<^isub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^isub>o(inv_bval'' b bv S)"
+lemma inv_bval''_correct: "s : \<gamma>\<^sub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^sub>o(inv_bval'' b bv S)"
proof(induction b arbitrary: S bv)
case Bc thus ?case by simp
next
@@ -159,24 +159,24 @@
subsubsection "Correctness"
-lemma step_step': "step (\<gamma>\<^isub>o S) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' S C)"
+lemma step_step': "step (\<gamma>\<^sub>o S) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' S C)"
unfolding step_def step'_def
by(rule gamma_Step_subcomm)
(auto simp: intro!: aval'_correct inv_bval''_correct in_gamma_update split: option.splits)
-lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
+lemma AI_correct: "AI c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
proof(simp add: CS_def AI_def)
assume 1: "pfp (step' \<top>) (bot c) = Some C"
have pfp': "step' \<top> C \<le> C" by(rule pfp_pfp[OF 1])
- have 2: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C" --"transfer the pfp'"
+ have 2: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C" --"transfer the pfp'"
proof(rule order_trans)
- show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)" by(rule step_step')
- show "... \<le> \<gamma>\<^isub>c C" by (metis mono_gamma_c[OF pfp'])
+ show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)" by(rule step_step')
+ show "... \<le> \<gamma>\<^sub>c C" by (metis mono_gamma_c[OF pfp'])
qed
- have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
- have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C"
- by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 2])
- thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
+ have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp[OF _ 1] step'_def)
+ have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
+ by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 2])
+ thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
qed
end
--- a/src/HOL/IMP/Abs_Int3.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int3.thy Tue Aug 13 16:25:47 2013 +0200
@@ -240,22 +240,22 @@
definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
"AI_wn c = pfp_wn (step' \<top>) (bot c)"
-lemma AI_wn_correct: "AI_wn c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c C"
+lemma AI_wn_correct: "AI_wn c = Some C \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c C"
proof(simp add: CS_def AI_wn_def)
assume 1: "pfp_wn (step' \<top>) (bot c) = Some C"
have 2: "strip C = c \<and> step' \<top> C \<le> C"
by(rule pfp_wn_pfp[where x="bot c"]) (simp_all add: 1 mono_step'_top)
- have pfp: "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c C"
+ have pfp: "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c C"
proof(rule order_trans)
- show "step (\<gamma>\<^isub>o \<top>) (\<gamma>\<^isub>c C) \<le> \<gamma>\<^isub>c (step' \<top> C)"
+ show "step (\<gamma>\<^sub>o \<top>) (\<gamma>\<^sub>c C) \<le> \<gamma>\<^sub>c (step' \<top> C)"
by(rule step_step')
- show "... \<le> \<gamma>\<^isub>c C"
+ show "... \<le> \<gamma>\<^sub>c C"
by(rule mono_gamma_c[OF conjunct2[OF 2]])
qed
- have 3: "strip (\<gamma>\<^isub>c C) = c" by(simp add: strip_pfp_wn[OF _ 1])
- have "lfp c (step (\<gamma>\<^isub>o \<top>)) \<le> \<gamma>\<^isub>c C"
- by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^isub>o \<top>)", OF 3 pfp])
- thus "lfp c (step UNIV) \<le> \<gamma>\<^isub>c C" by simp
+ have 3: "strip (\<gamma>\<^sub>c C) = c" by(simp add: strip_pfp_wn[OF _ 1])
+ have "lfp c (step (\<gamma>\<^sub>o \<top>)) \<le> \<gamma>\<^sub>c C"
+ by(rule lfp_lowerbound[simplified,where f="step (\<gamma>\<^sub>o \<top>)", OF 3 pfp])
+ thus "lfp c (step UNIV) \<le> \<gamma>\<^sub>c C" by simp
qed
end
@@ -405,8 +405,8 @@
done
-definition n_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^isub>s") where
-"n\<^isub>s S X = (\<Sum>x\<in>X. n(fun S x))"
+definition n_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>s") where
+"n\<^sub>s S X = (\<Sum>x\<in>X. n(fun S x))"
lemma n_s_narrow_rep:
assumes "finite X" "S1 = S2 on -X" "\<forall>x. S2 x \<le> S1 x" "\<forall>x. S1 x \<triangle> S2 x \<le> S1 x"
@@ -423,25 +423,25 @@
qed
lemma n_s_narrow: "finite X \<Longrightarrow> fun S1 = fun S2 on -X \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1
- \<Longrightarrow> n\<^isub>s (S1 \<triangle> S2) X < n\<^isub>s S1 X"
+ \<Longrightarrow> n\<^sub>s (S1 \<triangle> S2) X < n\<^sub>s S1 X"
apply(auto simp add: less_st_def n_s_def)
apply (transfer fixing: n)
apply(auto simp add: less_eq_st_rep_iff eq_st_def fun_eq_iff n_s_narrow_rep)
done
-definition n_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^isub>o") where
-"n\<^isub>o opt X = (case opt of None \<Rightarrow> 0 | Some S \<Rightarrow> n\<^isub>s S X + 1)"
+definition n_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>o") where
+"n\<^sub>o opt X = (case opt of None \<Rightarrow> 0 | Some S \<Rightarrow> n\<^sub>s S X + 1)"
lemma n_o_narrow:
"top_on_opt S1 (-X) \<Longrightarrow> top_on_opt S2 (-X) \<Longrightarrow> finite X
- \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1 \<Longrightarrow> n\<^isub>o (S1 \<triangle> S2) X < n\<^isub>o S1 X"
+ \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1 \<Longrightarrow> n\<^sub>o (S1 \<triangle> S2) X < n\<^sub>o S1 X"
apply(induction S1 S2 rule: narrow_option.induct)
apply(auto simp: n_o_def n_s_narrow)
done
-definition n_c :: "'av st option acom \<Rightarrow> nat" ("n\<^isub>c") where
-"n\<^isub>c C = listsum (map (\<lambda>a. n\<^isub>o a (vars C)) (annos C))"
+definition n_c :: "'av st option acom \<Rightarrow> nat" ("n\<^sub>c") where
+"n\<^sub>c C = listsum (map (\<lambda>a. n\<^sub>o a (vars C)) (annos C))"
lemma less_annos_iff: "(C1 < C2) = (C1 \<le> C2 \<and>
(\<exists>i<length (annos C1). annos C1 ! i < annos C2 ! i))"
@@ -449,7 +449,7 @@
lemma n_c_narrow: "strip C1 = strip C2
\<Longrightarrow> top_on_acom C1 (- vars C1) \<Longrightarrow> top_on_acom C2 (- vars C2)
- \<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n\<^isub>c (C1 \<triangle> C2) < n\<^isub>c C1"
+ \<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n\<^sub>c (C1 \<triangle> C2) < n\<^sub>c C1"
apply(auto simp: n_c_def narrow_acom_def listsum_setsum_nth)
apply(subgoal_tac "length(annos C2) = length(annos C1)")
prefer 2 apply (simp add: size_annos_same2)
--- a/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1.thy Tue Aug 13 16:25:47 2013 +0200
@@ -102,7 +102,7 @@
lemma in_rep_join_UpI: "s <:: S1 | s <:: S2 \<Longrightarrow> s <:: S1 \<squnion> S2"
by (metis in_rep_up_trans SL_top_class.join_ge1 SL_top_class.join_ge2)
-fun aval' :: "aexp \<Rightarrow> 'a astate up \<Rightarrow> 'a" ("aval\<^isup>#") where
+fun aval' :: "aexp \<Rightarrow> 'a astate up \<Rightarrow> 'a" ("aval\<^sup>#") where
"aval' _ bot = Bot" |
"aval' (N n) _ = num' n" |
"aval' (V x) (Up S) = lookup S x" |
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int0_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -203,26 +203,26 @@
qed
definition
- lpfp\<^isub>c :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" where
-"lpfp\<^isub>c f c = pfp f (\<bottom>\<^sub>c c)"
+ lpfp\<^sub>c :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom) \<Rightarrow> com \<Rightarrow> 'a option acom option" where
+"lpfp\<^sub>c f c = pfp f (\<bottom>\<^sub>c c)"
-lemma lpfpc_pfp: "lpfp\<^isub>c f c0 = Some c \<Longrightarrow> f c \<sqsubseteq> c"
-by(simp add: pfp_pfp lpfp\<^isub>c_def)
+lemma lpfpc_pfp: "lpfp\<^sub>c f c0 = Some c \<Longrightarrow> f c \<sqsubseteq> c"
+by(simp add: pfp_pfp lpfp\<^sub>c_def)
lemma strip_pfp:
assumes "\<And>x. g(f x) = g x" and "pfp f x0 = Some x" shows "g x = g x0"
using assms while_option_rule[where P = "%x. g x = g x0" and c = f]
unfolding pfp_def by metis
-lemma strip_lpfpc: assumes "\<And>c. strip(f c) = strip c" and "lpfp\<^isub>c f c = Some c'"
+lemma strip_lpfpc: assumes "\<And>c. strip(f c) = strip c" and "lpfp\<^sub>c f c = Some c'"
shows "strip c' = c"
-using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^isub>c_def]]
+using assms(1) strip_pfp[OF _ assms(2)[simplified lpfp\<^sub>c_def]]
by(metis strip_bot_acom)
lemma lpfpc_least:
assumes mono: "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
-and "strip p = c0" and "f p \<sqsubseteq> p" and lp: "lpfp\<^isub>c f c0 = Some c" shows "c \<sqsubseteq> p"
-using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^isub>c_def]]
+and "strip p = c0" and "f p \<sqsubseteq> p" and lp: "lpfp\<^sub>c f c0 = Some c" shows "c \<sqsubseteq> p"
+using pfp_least[OF _ _ bot_acom[OF `strip p = c0`] lp[simplified lpfp\<^sub>c_def]]
mono `f p \<sqsubseteq> p`
by blast
@@ -270,51 +270,51 @@
{S \<squnion> post c} WHILE b DO (step' Inv c) {Inv}"
definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
+"AI = lpfp\<^sub>c (step' \<top>)"
lemma strip_step'[simp]: "strip(step' S c) = strip c"
by(induct c arbitrary: S) (simp_all add: Let_def)
-abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>f == \<gamma>_fun \<gamma>"
+abbreviation \<gamma>\<^sub>f :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^sub>f == \<gamma>_fun \<gamma>"
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
+abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>f"
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
-lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
+lemma gamma_f_Top[simp]: "\<gamma>\<^sub>f Top = UNIV"
by(simp add: Top_fun_def \<gamma>_fun_def)
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
+lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o Top = UNIV"
by (simp add: Top_option_def)
(* FIXME (maybe also le \<rightarrow> sqle?) *)
-lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
+lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^sub>f f \<subseteq> \<gamma>\<^sub>f g"
by(auto simp: le_fun_def \<gamma>_fun_def dest: mono_gamma)
lemma mono_gamma_o:
- "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
+ "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^sub>o sa \<subseteq> \<gamma>\<^sub>o sa'"
by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
-lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
+lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^sub>c ca \<le> \<gamma>\<^sub>c ca'"
by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
text{* Soundness: *}
-lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+lemma aval'_sound: "s : \<gamma>\<^sub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
by (induct a) (auto simp: gamma_num' gamma_plus' \<gamma>_fun_def)
lemma in_gamma_update:
- "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(S(x := a))"
+ "\<lbrakk> s : \<gamma>\<^sub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>f(S(x := a))"
by(simp add: \<gamma>_fun_def)
lemma step_preserves_le:
- "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^isub>c (step' S' c')"
+ "\<lbrakk> S \<subseteq> \<gamma>\<^sub>o S'; c \<le> \<gamma>\<^sub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^sub>c (step' S' c')"
proof(induction c arbitrary: c' S S')
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
next
@@ -328,40 +328,40 @@
case (If b c1 c2 P)
then obtain c1' c2' P' where
"c' = IF b THEN c1' ELSE c2' {P'}"
- "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
+ "P \<subseteq> \<gamma>\<^sub>o P'" "c1 \<le> \<gamma>\<^sub>c c1'" "c2 \<le> \<gamma>\<^sub>c c2'"
by (fastforce simp: If_le map_acom_If)
- moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
- by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
- moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
- by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
- ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
+ moreover have "post c1 \<subseteq> \<gamma>\<^sub>o(post c1' \<squnion> post c2')"
+ by (metis (no_types) `c1 \<le> \<gamma>\<^sub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+ moreover have "post c2 \<subseteq> \<gamma>\<^sub>o(post c1' \<squnion> post c2')"
+ by (metis (no_types) `c2 \<le> \<gamma>\<^sub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+ ultimately show ?case using `S \<subseteq> \<gamma>\<^sub>o S'` by (simp add: If.IH subset_iff)
next
case (While I b c1 P)
then obtain c1' I' P' where
"c' = {I'} WHILE b DO c1' {P'}"
- "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
+ "I \<subseteq> \<gamma>\<^sub>o I'" "P \<subseteq> \<gamma>\<^sub>o P'" "c1 \<le> \<gamma>\<^sub>c c1'"
by (fastforce simp: map_acom_While While_le)
- moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
- using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
+ moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^sub>o (S' \<squnion> post c1')"
+ using `S \<subseteq> \<gamma>\<^sub>o S'` le_post[OF `c1 \<le> \<gamma>\<^sub>c c1'`, simplified]
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
ultimately show ?case by (simp add: While.IH subset_iff)
qed
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c c'"
proof(simp add: CS_def AI_def)
- assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+ assume 1: "lpfp\<^sub>c (step' \<top>) c = Some c'"
have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
- have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+ have 3: "strip (\<gamma>\<^sub>c (step' \<top> c')) = c"
by(simp add: strip_lpfpc[OF _ 1])
- have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ have "lfp (step UNIV) c \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule lfp_lowerbound[simplified,OF 3])
- show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ show "step UNIV (\<gamma>\<^sub>c (step' \<top> c')) \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule step_preserves_le[OF _ _])
- show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
- show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+ show "UNIV \<subseteq> \<gamma>\<^sub>o \<top>" by simp
+ show "\<gamma>\<^sub>c (step' \<top> c') \<le> \<gamma>\<^sub>c c'" by(rule mono_gamma_c[OF 2])
qed
qed
- with 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+ with 2 show "lfp (step UNIV) c \<le> \<gamma>\<^sub>c c'"
by (blast intro: mono_gamma_c order_trans)
qed
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -17,7 +17,7 @@
"aval' (V x) S = lookup S x" |
"aval' (Plus a1 a2) S = plus' (aval' a1 S) (aval' a2 S)"
-lemma aval'_sound: "s : \<gamma>\<^isub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
+lemma aval'_sound: "s : \<gamma>\<^sub>f S \<Longrightarrow> aval a s : \<gamma>(aval' a S)"
by (induction a) (auto simp: gamma_num' gamma_plus' \<gamma>_st_def lookup_def)
end
@@ -41,7 +41,7 @@
{S \<squnion> post c} WHILE b DO step' Inv c {Inv}"
definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
+"AI = lpfp\<^sub>c (step' \<top>)"
lemma strip_step'[simp]: "strip(step' S c) = strip c"
@@ -51,14 +51,14 @@
text{* Soundness: *}
lemma in_gamma_update:
- "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
+ "\<lbrakk> s : \<gamma>\<^sub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>f(update S x a)"
by(simp add: \<gamma>_st_def lookup_update)
text{* The soundness proofs are textually identical to the ones for the step
function operating on states as functions. *}
lemma step_preserves_le:
- "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; c \<le> \<gamma>\<^isub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^isub>c (step' S' c')"
+ "\<lbrakk> S \<subseteq> \<gamma>\<^sub>o S'; c \<le> \<gamma>\<^sub>c c' \<rbrakk> \<Longrightarrow> step S c \<le> \<gamma>\<^sub>c (step' S' c')"
proof(induction c arbitrary: c' S S')
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
next
@@ -72,40 +72,40 @@
case (If b c1 c2 P)
then obtain c1' c2' P' where
"c' = IF b THEN c1' ELSE c2' {P'}"
- "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'" "c2 \<le> \<gamma>\<^isub>c c2'"
+ "P \<subseteq> \<gamma>\<^sub>o P'" "c1 \<le> \<gamma>\<^sub>c c1'" "c2 \<le> \<gamma>\<^sub>c c2'"
by (fastforce simp: If_le map_acom_If)
- moreover have "post c1 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
- by (metis (no_types) `c1 \<le> \<gamma>\<^isub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
- moreover have "post c2 \<subseteq> \<gamma>\<^isub>o(post c1' \<squnion> post c2')"
- by (metis (no_types) `c2 \<le> \<gamma>\<^isub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
- ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'` by (simp add: If.IH subset_iff)
+ moreover have "post c1 \<subseteq> \<gamma>\<^sub>o(post c1' \<squnion> post c2')"
+ by (metis (no_types) `c1 \<le> \<gamma>\<^sub>c c1'` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+ moreover have "post c2 \<subseteq> \<gamma>\<^sub>o(post c1' \<squnion> post c2')"
+ by (metis (no_types) `c2 \<le> \<gamma>\<^sub>c c2'` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+ ultimately show ?case using `S \<subseteq> \<gamma>\<^sub>o S'` by (simp add: If.IH subset_iff)
next
case (While I b c1 P)
then obtain c1' I' P' where
"c' = {I'} WHILE b DO c1' {P'}"
- "I \<subseteq> \<gamma>\<^isub>o I'" "P \<subseteq> \<gamma>\<^isub>o P'" "c1 \<le> \<gamma>\<^isub>c c1'"
+ "I \<subseteq> \<gamma>\<^sub>o I'" "P \<subseteq> \<gamma>\<^sub>o P'" "c1 \<le> \<gamma>\<^sub>c c1'"
by (fastforce simp: map_acom_While While_le)
- moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post c1')"
- using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `c1 \<le> \<gamma>\<^isub>c c1'`, simplified]
+ moreover have "S \<union> post c1 \<subseteq> \<gamma>\<^sub>o (S' \<squnion> post c1')"
+ using `S \<subseteq> \<gamma>\<^sub>o S'` le_post[OF `c1 \<le> \<gamma>\<^sub>c c1'`, simplified]
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
ultimately show ?case by (simp add: While.IH subset_iff)
qed
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c c'"
proof(simp add: CS_def AI_def)
- assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+ assume 1: "lpfp\<^sub>c (step' \<top>) c = Some c'"
have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
- have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+ have 3: "strip (\<gamma>\<^sub>c (step' \<top> c')) = c"
by(simp add: strip_lpfpc[OF _ 1])
- have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ have "lfp (step UNIV) c \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule lfp_lowerbound[simplified,OF 3])
- show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ show "step UNIV (\<gamma>\<^sub>c (step' \<top> c')) \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule step_preserves_le[OF _ _])
- show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
- show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+ show "UNIV \<subseteq> \<gamma>\<^sub>o \<top>" by simp
+ show "\<gamma>\<^sub>c (step' \<top> c') \<le> \<gamma>\<^sub>c c'" by(rule mono_gamma_c[OF 2])
qed
qed
- from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+ from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^sub>c c'"
by (blast intro: mono_gamma_c order_trans)
qed
@@ -367,8 +367,8 @@
fixes f :: "(('a::SL_top)option acom \<Rightarrow> 'a option acom)"
assumes "acc {(x::'a,y). x \<sqsubseteq> y}" and "\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y"
and "\<And>c. strip(f c) = strip c"
- shows "\<exists>c'. lpfp\<^isub>c f c = Some c'"
-unfolding lpfp\<^isub>c_def
+ shows "\<exists>c'. lpfp\<^sub>c f c = Some c'"
+unfolding lpfp\<^sub>c_def
apply(rule pfp_termination)
apply(erule assms(2))
apply(rule acc_acom[OF acc_option[OF assms(1)]])
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -70,14 +70,14 @@
Val_abs1 where \<gamma> = \<gamma> for \<gamma> :: "'av::L_top_bot \<Rightarrow> val set"
begin
-lemma in_gamma_join_UpI: "s : \<gamma>\<^isub>o S1 \<or> s : \<gamma>\<^isub>o S2 \<Longrightarrow> s : \<gamma>\<^isub>o(S1 \<squnion> S2)"
+lemma in_gamma_join_UpI: "s : \<gamma>\<^sub>o S1 \<or> s : \<gamma>\<^sub>o S2 \<Longrightarrow> s : \<gamma>\<^sub>o(S1 \<squnion> S2)"
by (metis (no_types) join_ge1 join_ge2 mono_gamma_o set_rev_mp)
fun aval'' :: "aexp \<Rightarrow> 'av st option \<Rightarrow> 'av" where
"aval'' e None = \<bottom>" |
"aval'' e (Some sa) = aval' e sa"
-lemma aval''_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
+lemma aval''_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval a s : \<gamma>(aval'' a S)"
by(cases S)(simp add: aval'_sound)+
subsubsection "Backward analysis"
@@ -111,12 +111,12 @@
(let (res1,res2) = filter_less' res (aval'' e1 S) (aval'' e2 S)
in afilter e1 res1 (afilter e2 res2 S))"
-lemma afilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^isub>o (afilter e a S)"
+lemma afilter_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> aval e s : \<gamma> a \<Longrightarrow> s : \<gamma>\<^sub>o (afilter e a S)"
proof(induction e arbitrary: a S)
case N thus ?case by simp (metis test_num')
next
case (V x)
- obtain S' where "S = Some S'" and "s : \<gamma>\<^isub>f S'" using `s : \<gamma>\<^isub>o S`
+ obtain S' where "S = Some S'" and "s : \<gamma>\<^sub>f S'" using `s : \<gamma>\<^sub>o S`
by(auto simp: in_gamma_option_iff)
moreover hence "s x : \<gamma> (lookup S' x)" by(simp add: \<gamma>_st_def)
moreover have "s x : \<gamma> a" using V by simp
@@ -129,7 +129,7 @@
by (auto split: prod.split)
qed
-lemma bfilter_sound: "s : \<gamma>\<^isub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^isub>o(bfilter b bv S)"
+lemma bfilter_sound: "s : \<gamma>\<^sub>o S \<Longrightarrow> bv = bval b s \<Longrightarrow> s : \<gamma>\<^sub>o(bfilter b bv S)"
proof(induction b arbitrary: S bv)
case Bc thus ?case by simp
next
@@ -158,7 +158,7 @@
{bfilter b False Inv}"
definition AI :: "com \<Rightarrow> 'av st option acom option" where
-"AI = lpfp\<^isub>c (step' \<top>)"
+"AI = lpfp\<^sub>c (step' \<top>)"
lemma strip_step'[simp]: "strip(step' S c) = strip c"
by(induct c arbitrary: S) (simp_all add: Let_def)
@@ -167,11 +167,11 @@
subsubsection "Soundness"
lemma in_gamma_update:
- "\<lbrakk> s : \<gamma>\<^isub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^isub>f(update S x a)"
+ "\<lbrakk> s : \<gamma>\<^sub>f S; i : \<gamma> a \<rbrakk> \<Longrightarrow> s(x := i) : \<gamma>\<^sub>f(update S x a)"
by(simp add: \<gamma>_st_def lookup_update)
lemma step_preserves_le:
- "\<lbrakk> S \<subseteq> \<gamma>\<^isub>o S'; cs \<le> \<gamma>\<^isub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^isub>c (step' S' ca)"
+ "\<lbrakk> S \<subseteq> \<gamma>\<^sub>o S'; cs \<le> \<gamma>\<^sub>c ca \<rbrakk> \<Longrightarrow> step S cs \<le> \<gamma>\<^sub>c (step' S' ca)"
proof(induction cs arbitrary: ca S S')
case SKIP thus ?case by(auto simp:SKIP_le map_acom_SKIP)
next
@@ -185,41 +185,41 @@
case (If b cs1 cs2 P)
then obtain ca1 ca2 Pa where
"ca= IF b THEN ca1 ELSE ca2 {Pa}"
- "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1" "cs2 \<le> \<gamma>\<^isub>c ca2"
+ "P \<subseteq> \<gamma>\<^sub>o Pa" "cs1 \<le> \<gamma>\<^sub>c ca1" "cs2 \<le> \<gamma>\<^sub>c ca2"
by (fastforce simp: If_le map_acom_If)
- moreover have "post cs1 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
- by (metis (no_types) `cs1 \<le> \<gamma>\<^isub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
- moreover have "post cs2 \<subseteq> \<gamma>\<^isub>o(post ca1 \<squnion> post ca2)"
- by (metis (no_types) `cs2 \<le> \<gamma>\<^isub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
- ultimately show ?case using `S \<subseteq> \<gamma>\<^isub>o S'`
+ moreover have "post cs1 \<subseteq> \<gamma>\<^sub>o(post ca1 \<squnion> post ca2)"
+ by (metis (no_types) `cs1 \<le> \<gamma>\<^sub>c ca1` join_ge1 le_post mono_gamma_o order_trans post_map_acom)
+ moreover have "post cs2 \<subseteq> \<gamma>\<^sub>o(post ca1 \<squnion> post ca2)"
+ by (metis (no_types) `cs2 \<le> \<gamma>\<^sub>c ca2` join_ge2 le_post mono_gamma_o order_trans post_map_acom)
+ ultimately show ?case using `S \<subseteq> \<gamma>\<^sub>o S'`
by (simp add: If.IH subset_iff bfilter_sound)
next
case (While I b cs1 P)
then obtain ca1 Ia Pa where
"ca = {Ia} WHILE b DO ca1 {Pa}"
- "I \<subseteq> \<gamma>\<^isub>o Ia" "P \<subseteq> \<gamma>\<^isub>o Pa" "cs1 \<le> \<gamma>\<^isub>c ca1"
+ "I \<subseteq> \<gamma>\<^sub>o Ia" "P \<subseteq> \<gamma>\<^sub>o Pa" "cs1 \<le> \<gamma>\<^sub>c ca1"
by (fastforce simp: map_acom_While While_le)
- moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^isub>o (S' \<squnion> post ca1)"
- using `S \<subseteq> \<gamma>\<^isub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^isub>c ca1`, simplified]
+ moreover have "S \<union> post cs1 \<subseteq> \<gamma>\<^sub>o (S' \<squnion> post ca1)"
+ using `S \<subseteq> \<gamma>\<^sub>o S'` le_post[OF `cs1 \<le> \<gamma>\<^sub>c ca1`, simplified]
by (metis (no_types) join_ge1 join_ge2 le_sup_iff mono_gamma_o order_trans)
ultimately show ?case by (simp add: While.IH subset_iff bfilter_sound)
qed
-lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+lemma AI_sound: "AI c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c c'"
proof(simp add: CS_def AI_def)
- assume 1: "lpfp\<^isub>c (step' \<top>) c = Some c'"
+ assume 1: "lpfp\<^sub>c (step' \<top>) c = Some c'"
have 2: "step' \<top> c' \<sqsubseteq> c'" by(rule lpfpc_pfp[OF 1])
- have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c"
+ have 3: "strip (\<gamma>\<^sub>c (step' \<top> c')) = c"
by(simp add: strip_lpfpc[OF _ 1])
- have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ have "lfp (step UNIV) c \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule lfp_lowerbound[simplified,OF 3])
- show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ show "step UNIV (\<gamma>\<^sub>c (step' \<top> c')) \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule step_preserves_le[OF _ _])
- show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
- show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+ show "UNIV \<subseteq> \<gamma>\<^sub>o \<top>" by simp
+ show "\<gamma>\<^sub>c (step' \<top> c') \<le> \<gamma>\<^sub>c c'" by(rule mono_gamma_c[OF 2])
qed
qed
- from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+ from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^sub>c c'"
by (blast intro: mono_gamma_c order_trans)
qed
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -205,21 +205,21 @@
definition AI_wn :: "com \<Rightarrow> 'av st option acom option" where
"AI_wn = pfp_wn (step' \<top>)"
-lemma AI_wn_sound: "AI_wn c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^isub>c c'"
+lemma AI_wn_sound: "AI_wn c = Some c' \<Longrightarrow> CS c \<le> \<gamma>\<^sub>c c'"
proof(simp add: CS_def AI_wn_def)
assume 1: "pfp_wn (step' \<top>) c = Some c'"
from pfp_wn_pfp[OF mono_step'2 1]
have 2: "step' \<top> c' \<sqsubseteq> c'" .
- have 3: "strip (\<gamma>\<^isub>c (step' \<top> c')) = c" by(simp add: strip_pfp_wn[OF _ 1])
- have "lfp (step UNIV) c \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ have 3: "strip (\<gamma>\<^sub>c (step' \<top> c')) = c" by(simp add: strip_pfp_wn[OF _ 1])
+ have "lfp (step UNIV) c \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule lfp_lowerbound[simplified,OF 3])
- show "step UNIV (\<gamma>\<^isub>c (step' \<top> c')) \<le> \<gamma>\<^isub>c (step' \<top> c')"
+ show "step UNIV (\<gamma>\<^sub>c (step' \<top> c')) \<le> \<gamma>\<^sub>c (step' \<top> c')"
proof(rule step_preserves_le[OF _ _])
- show "UNIV \<subseteq> \<gamma>\<^isub>o \<top>" by simp
- show "\<gamma>\<^isub>c (step' \<top> c') \<le> \<gamma>\<^isub>c c'" by(rule mono_gamma_c[OF 2])
+ show "UNIV \<subseteq> \<gamma>\<^sub>o \<top>" by simp
+ show "\<gamma>\<^sub>c (step' \<top> c') \<le> \<gamma>\<^sub>c c'" by(rule mono_gamma_c[OF 2])
qed
qed
- from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^isub>c c'"
+ from this 2 show "lfp (step UNIV) c \<le> \<gamma>\<^sub>c c'"
by (blast intro: mono_gamma_c order_trans)
qed
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_State_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -61,32 +61,32 @@
locale Gamma = Val_abs where \<gamma>=\<gamma> for \<gamma> :: "'av::SL_top \<Rightarrow> val set"
begin
-abbreviation \<gamma>\<^isub>f :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>f == \<gamma>_st \<gamma>"
+abbreviation \<gamma>\<^sub>f :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^sub>f == \<gamma>_st \<gamma>"
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>f"
+abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>f"
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
-lemma gamma_f_Top[simp]: "\<gamma>\<^isub>f Top = UNIV"
+lemma gamma_f_Top[simp]: "\<gamma>\<^sub>f Top = UNIV"
by(auto simp: Top_st_def \<gamma>_st_def lookup_def)
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o Top = UNIV"
+lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o Top = UNIV"
by (simp add: Top_option_def)
(* FIXME (maybe also le \<rightarrow> sqle?) *)
-lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^isub>f f \<subseteq> \<gamma>\<^isub>f g"
+lemma mono_gamma_f: "f \<sqsubseteq> g \<Longrightarrow> \<gamma>\<^sub>f f \<subseteq> \<gamma>\<^sub>f g"
apply(simp add:\<gamma>_st_def subset_iff lookup_def le_st_def split: if_splits)
by (metis UNIV_I mono_gamma gamma_Top subsetD)
lemma mono_gamma_o:
- "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^isub>o sa \<subseteq> \<gamma>\<^isub>o sa'"
+ "sa \<sqsubseteq> sa' \<Longrightarrow> \<gamma>\<^sub>o sa \<subseteq> \<gamma>\<^sub>o sa'"
by(induction sa sa' rule: le_option.induct)(simp_all add: mono_gamma_f)
-lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^isub>c ca \<le> \<gamma>\<^isub>c ca'"
+lemma mono_gamma_c: "ca \<sqsubseteq> ca' \<Longrightarrow> \<gamma>\<^sub>c ca \<le> \<gamma>\<^sub>c ca'"
by (induction ca ca' rule: le_acom.induct) (simp_all add:mono_gamma_o)
lemma in_gamma_option_iff:
--- a/src/HOL/IMP/Abs_Int_ITP/Collecting_ITP.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Collecting_ITP.thy Tue Aug 13 16:25:47 2013 +0200
@@ -63,14 +63,14 @@
end
-fun sub\<^isub>1 :: "'a acom \<Rightarrow> 'a acom" where
-"sub\<^isub>1(c1;;c2) = c1" |
-"sub\<^isub>1(IF b THEN c1 ELSE c2 {S}) = c1" |
-"sub\<^isub>1({I} WHILE b DO c {P}) = c"
+fun sub\<^sub>1 :: "'a acom \<Rightarrow> 'a acom" where
+"sub\<^sub>1(c1;;c2) = c1" |
+"sub\<^sub>1(IF b THEN c1 ELSE c2 {S}) = c1" |
+"sub\<^sub>1({I} WHILE b DO c {P}) = c"
-fun sub\<^isub>2 :: "'a acom \<Rightarrow> 'a acom" where
-"sub\<^isub>2(c1;;c2) = c2" |
-"sub\<^isub>2(IF b THEN c1 ELSE c2 {S}) = c2"
+fun sub\<^sub>2 :: "'a acom \<Rightarrow> 'a acom" where
+"sub\<^sub>2(c1;;c2) = c2" |
+"sub\<^sub>2(IF b THEN c1 ELSE c2 {S}) = c2"
fun invar :: "'a acom \<Rightarrow> 'a" where
"invar({I} WHILE b DO c {P}) = I"
@@ -80,13 +80,13 @@
"lift F com.SKIP M = (SKIP {F (post ` M)})" |
"lift F (x ::= a) M = (x ::= a {F (post ` M)})" |
"lift F (c1;;c2) M =
- lift F c1 (sub\<^isub>1 ` M);; lift F c2 (sub\<^isub>2 ` M)" |
+ lift F c1 (sub\<^sub>1 ` M);; lift F c2 (sub\<^sub>2 ` M)" |
"lift F (IF b THEN c1 ELSE c2) M =
- IF b THEN lift F c1 (sub\<^isub>1 ` M) ELSE lift F c2 (sub\<^isub>2 ` M)
+ IF b THEN lift F c1 (sub\<^sub>1 ` M) ELSE lift F c2 (sub\<^sub>2 ` M)
{F (post ` M)}" |
"lift F (WHILE b DO c) M =
{F (invar ` M)}
- WHILE b DO lift F c (sub\<^isub>1 ` M)
+ WHILE b DO lift F c (sub\<^sub>1 ` M)
{F (post ` M)}"
interpretation Complete_Lattice_ix "%c. {c'. strip c' = c}" "lift Inter"
--- a/src/HOL/IMP/Abs_State.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Abs_State.thy Tue Aug 13 16:25:47 2013 +0200
@@ -132,29 +132,29 @@
for \<gamma> :: "'av::semilattice_sup_top \<Rightarrow> val set"
begin
-abbreviation \<gamma>\<^isub>s :: "'av st \<Rightarrow> state set"
-where "\<gamma>\<^isub>s == \<gamma>_st \<gamma>"
+abbreviation \<gamma>\<^sub>s :: "'av st \<Rightarrow> state set"
+where "\<gamma>\<^sub>s == \<gamma>_st \<gamma>"
-abbreviation \<gamma>\<^isub>o :: "'av st option \<Rightarrow> state set"
-where "\<gamma>\<^isub>o == \<gamma>_option \<gamma>\<^isub>s"
+abbreviation \<gamma>\<^sub>o :: "'av st option \<Rightarrow> state set"
+where "\<gamma>\<^sub>o == \<gamma>_option \<gamma>\<^sub>s"
-abbreviation \<gamma>\<^isub>c :: "'av st option acom \<Rightarrow> state set acom"
-where "\<gamma>\<^isub>c == map_acom \<gamma>\<^isub>o"
+abbreviation \<gamma>\<^sub>c :: "'av st option acom \<Rightarrow> state set acom"
+where "\<gamma>\<^sub>c == map_acom \<gamma>\<^sub>o"
-lemma gamma_s_top[simp]: "\<gamma>\<^isub>s \<top> = UNIV"
+lemma gamma_s_top[simp]: "\<gamma>\<^sub>s \<top> = UNIV"
by(auto simp: \<gamma>_st_def fun_top)
-lemma gamma_o_Top[simp]: "\<gamma>\<^isub>o \<top> = UNIV"
+lemma gamma_o_Top[simp]: "\<gamma>\<^sub>o \<top> = UNIV"
by (simp add: top_option_def)
-lemma mono_gamma_s: "f \<le> g \<Longrightarrow> \<gamma>\<^isub>s f \<subseteq> \<gamma>\<^isub>s g"
+lemma mono_gamma_s: "f \<le> g \<Longrightarrow> \<gamma>\<^sub>s f \<subseteq> \<gamma>\<^sub>s g"
by(simp add:\<gamma>_st_def le_st_iff subset_iff) (metis mono_gamma subsetD)
lemma mono_gamma_o:
- "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^isub>o S1 \<subseteq> \<gamma>\<^isub>o S2"
+ "S1 \<le> S2 \<Longrightarrow> \<gamma>\<^sub>o S1 \<subseteq> \<gamma>\<^sub>o S2"
by(induction S1 S2 rule: less_eq_option.induct)(simp_all add: mono_gamma_s)
-lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^isub>c C1 \<le> \<gamma>\<^isub>c C2"
+lemma mono_gamma_c: "C1 \<le> C2 \<Longrightarrow> \<gamma>\<^sub>c C1 \<le> \<gamma>\<^sub>c C2"
by (simp add: less_eq_acom_def mono_gamma_o size_annos anno_map_acom size_annos_same[of C1 C2])
lemma in_gamma_option_iff:
--- a/src/HOL/IMP/BExp.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/BExp.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,8 +10,8 @@
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where
"bval (Bc v) s = v" |
"bval (Not b) s = (\<not> bval b s)" |
-"bval (And b\<^isub>1 b\<^isub>2) s = (bval b\<^isub>1 s \<and> bval b\<^isub>2 s)" |
-"bval (Less a\<^isub>1 a\<^isub>2) s = (aval a\<^isub>1 s < aval a\<^isub>2 s)"
+"bval (And b\<^sub>1 b\<^sub>2) s = (bval b\<^sub>1 s \<and> bval b\<^sub>2 s)" |
+"bval (Less a\<^sub>1 a\<^sub>2) s = (aval a\<^sub>1 s < aval a\<^sub>2 s)"
text_raw{*}%endsnip*}
value "bval (Less (V ''x'') (Plus (N 3) (V ''y'')))
@@ -33,8 +33,8 @@
text_raw{*\snip{BExplessdef}{0}{2}{% *}
fun less :: "aexp \<Rightarrow> aexp \<Rightarrow> bexp" where
-"less (N n\<^isub>1) (N n\<^isub>2) = Bc(n\<^isub>1 < n\<^isub>2)" |
-"less a\<^isub>1 a\<^isub>2 = Less a\<^isub>1 a\<^isub>2"
+"less (N n\<^sub>1) (N n\<^sub>2) = Bc(n\<^sub>1 < n\<^sub>2)" |
+"less a\<^sub>1 a\<^sub>2 = Less a\<^sub>1 a\<^sub>2"
text_raw{*}%endsnip*}
lemma [simp]: "bval (less a1 a2) s = (aval a1 s < aval a2 s)"
@@ -48,7 +48,7 @@
"and b (Bc True) = b" |
"and (Bc False) b = Bc False" |
"and b (Bc False) = Bc False" |
-"and b\<^isub>1 b\<^isub>2 = And b\<^isub>1 b\<^isub>2"
+"and b\<^sub>1 b\<^sub>2 = And b\<^sub>1 b\<^sub>2"
text_raw{*}%endsnip*}
lemma bval_and[simp]: "bval (and b1 b2) s = (bval b1 s \<and> bval b2 s)"
@@ -74,8 +74,8 @@
fun bsimp :: "bexp \<Rightarrow> bexp" where
"bsimp (Bc v) = Bc v" |
"bsimp (Not b) = not(bsimp b)" |
-"bsimp (And b\<^isub>1 b\<^isub>2) = and (bsimp b\<^isub>1) (bsimp b\<^isub>2)" |
-"bsimp (Less a\<^isub>1 a\<^isub>2) = less (asimp a\<^isub>1) (asimp a\<^isub>2)"
+"bsimp (And b\<^sub>1 b\<^sub>2) = and (bsimp b\<^sub>1) (bsimp b\<^sub>2)" |
+"bsimp (Less a\<^sub>1 a\<^sub>2) = less (asimp a\<^sub>1) (asimp a\<^sub>2)"
text_raw{*}%endsnip*}
value "bsimp (And (Less (N 0) (N 1)) b)"
--- a/src/HOL/IMP/Big_Step.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Big_Step.thy Tue Aug 13 16:25:47 2013 +0200
@@ -16,13 +16,13 @@
where
Skip: "(SKIP,s) \<Rightarrow> s" |
Assign: "(x ::= a,s) \<Rightarrow> s(x := aval a s)" |
-Seq: "\<lbrakk> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2; (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow> (c\<^isub>1;;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
-IfTrue: "\<lbrakk> bval b s; (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
-IfFalse: "\<lbrakk> \<not>bval b s; (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
+Seq: "\<lbrakk> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
+IfTrue: "\<lbrakk> bval b s; (c\<^sub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
+IfFalse: "\<lbrakk> \<not>bval b s; (c\<^sub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s) \<Rightarrow> s" |
WhileTrue:
-"\<lbrakk> bval b s\<^isub>1; (c,s\<^isub>1) \<Rightarrow> s\<^isub>2; (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk>
-\<Longrightarrow> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3"
+"\<lbrakk> bval b s\<^sub>1; (c,s\<^sub>1) \<Rightarrow> s\<^sub>2; (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk>
+\<Longrightarrow> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3"
text_raw{*}%endsnip*}
text_raw{*\snip{BigStepEx}{1}{2}{% *}
@@ -283,19 +283,19 @@
"(c,s) \<Rightarrow> t \<Longrightarrow> (c,s) \<Rightarrow> t' \<Longrightarrow> t' = t"
proof (induction arbitrary: t' rule: big_step.induct)
-- "the only interesting case, @{text WhileTrue}:"
- fix b c s s\<^isub>1 t t'
+ fix b c s s\<^sub>1 t t'
-- "The assumptions of the rule:"
- assume "bval b s" and "(c,s) \<Rightarrow> s\<^isub>1" and "(WHILE b DO c,s\<^isub>1) \<Rightarrow> t"
+ assume "bval b s" and "(c,s) \<Rightarrow> s\<^sub>1" and "(WHILE b DO c,s\<^sub>1) \<Rightarrow> t"
-- {* Ind.Hyp; note the @{text"\<And>"} because of arbitrary: *}
- assume IHc: "\<And>t'. (c,s) \<Rightarrow> t' \<Longrightarrow> t' = s\<^isub>1"
- assume IHw: "\<And>t'. (WHILE b DO c,s\<^isub>1) \<Rightarrow> t' \<Longrightarrow> t' = t"
+ assume IHc: "\<And>t'. (c,s) \<Rightarrow> t' \<Longrightarrow> t' = s\<^sub>1"
+ assume IHw: "\<And>t'. (WHILE b DO c,s\<^sub>1) \<Rightarrow> t' \<Longrightarrow> t' = t"
-- "Premise of implication:"
assume "(WHILE b DO c,s) \<Rightarrow> t'"
- with `bval b s` obtain s\<^isub>1' where
- c: "(c,s) \<Rightarrow> s\<^isub>1'" and
- w: "(WHILE b DO c,s\<^isub>1') \<Rightarrow> t'"
+ with `bval b s` obtain s\<^sub>1' where
+ c: "(c,s) \<Rightarrow> s\<^sub>1'" and
+ w: "(WHILE b DO c,s\<^sub>1') \<Rightarrow> t'"
by auto
- from c IHc have "s\<^isub>1' = s\<^isub>1" by blast
+ from c IHc have "s\<^sub>1' = s\<^sub>1" by blast
with w IHw show "t' = t" by blast
qed blast+ -- "prove the rest automatically"
text_raw{*}%endsnip*}
--- a/src/HOL/IMP/C_like.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/C_like.thy Tue Aug 13 16:25:47 2013 +0200
@@ -9,15 +9,15 @@
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> nat" where
"aval (N n) s = n" |
"aval (!a) s = s(aval a s)" |
-"aval (Plus a\<^isub>1 a\<^isub>2) s = aval a\<^isub>1 s + aval a\<^isub>2 s"
+"aval (Plus a\<^sub>1 a\<^sub>2) s = aval a\<^sub>1 s + aval a\<^sub>2 s"
datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp
primrec bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where
"bval (Bc v) _ = v" |
"bval (Not b) s = (\<not> bval b s)" |
-"bval (And b\<^isub>1 b\<^isub>2) s = (if bval b\<^isub>1 s then bval b\<^isub>2 s else False)" |
-"bval (Less a\<^isub>1 a\<^isub>2) s = (aval a\<^isub>1 s < aval a\<^isub>2 s)"
+"bval (And b\<^sub>1 b\<^sub>2) s = (if bval b\<^sub>1 s then bval b\<^sub>2 s else False)" |
+"bval (Less a\<^sub>1 a\<^sub>2) s = (aval a\<^sub>1 s < aval a\<^sub>2 s)"
datatype
@@ -34,18 +34,18 @@
Skip: "(SKIP,sn) \<Rightarrow> sn" |
Assign: "(lhs ::= a,s,n) \<Rightarrow> (s(aval lhs s := aval a s),n)" |
New: "(New lhs a,s,n) \<Rightarrow> (s(aval lhs s := n), n+aval a s)" |
-Seq: "\<lbrakk> (c\<^isub>1,sn\<^isub>1) \<Rightarrow> sn\<^isub>2; (c\<^isub>2,sn\<^isub>2) \<Rightarrow> sn\<^isub>3 \<rbrakk> \<Longrightarrow>
- (c\<^isub>1;c\<^isub>2, sn\<^isub>1) \<Rightarrow> sn\<^isub>3" |
+Seq: "\<lbrakk> (c\<^sub>1,sn\<^sub>1) \<Rightarrow> sn\<^sub>2; (c\<^sub>2,sn\<^sub>2) \<Rightarrow> sn\<^sub>3 \<rbrakk> \<Longrightarrow>
+ (c\<^sub>1;c\<^sub>2, sn\<^sub>1) \<Rightarrow> sn\<^sub>3" |
-IfTrue: "\<lbrakk> bval b s; (c\<^isub>1,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow>
- (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s,n) \<Rightarrow> tn" |
-IfFalse: "\<lbrakk> \<not>bval b s; (c\<^isub>2,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow>
- (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s,n) \<Rightarrow> tn" |
+IfTrue: "\<lbrakk> bval b s; (c\<^sub>1,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow>
+ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s,n) \<Rightarrow> tn" |
+IfFalse: "\<lbrakk> \<not>bval b s; (c\<^sub>2,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow>
+ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s,n) \<Rightarrow> tn" |
WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s,n) \<Rightarrow> (s,n)" |
WhileTrue:
- "\<lbrakk> bval b s\<^isub>1; (c,s\<^isub>1,n) \<Rightarrow> sn\<^isub>2; (WHILE b DO c, sn\<^isub>2) \<Rightarrow> sn\<^isub>3 \<rbrakk> \<Longrightarrow>
- (WHILE b DO c, s\<^isub>1,n) \<Rightarrow> sn\<^isub>3"
+ "\<lbrakk> bval b s\<^sub>1; (c,s\<^sub>1,n) \<Rightarrow> sn\<^sub>2; (WHILE b DO c, sn\<^sub>2) \<Rightarrow> sn\<^sub>3 \<rbrakk> \<Longrightarrow>
+ (WHILE b DO c, s\<^sub>1,n) \<Rightarrow> sn\<^sub>3"
code_pred big_step .
--- a/src/HOL/IMP/Compiler.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Compiler.thy Tue Aug 13 16:25:47 2013 +0200
@@ -201,10 +201,10 @@
fun ccomp :: "com \<Rightarrow> instr list" where
"ccomp SKIP = []" |
"ccomp (x ::= a) = acomp a @ [STORE x]" |
-"ccomp (c\<^isub>1;;c\<^isub>2) = ccomp c\<^isub>1 @ ccomp c\<^isub>2" |
-"ccomp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) =
- (let cc\<^isub>1 = ccomp c\<^isub>1; cc\<^isub>2 = ccomp c\<^isub>2; cb = bcomp b False (size cc\<^isub>1 + 1)
- in cb @ cc\<^isub>1 @ JMP (size cc\<^isub>2) # cc\<^isub>2)" |
+"ccomp (c\<^sub>1;;c\<^sub>2) = ccomp c\<^sub>1 @ ccomp c\<^sub>2" |
+"ccomp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) =
+ (let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2; cb = bcomp b False (size cc\<^sub>1 + 1)
+ in cb @ cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" |
"ccomp (WHILE b DO c) =
(let cc = ccomp c; cb = bcomp b False (size cc + 1)
in cb @ cc @ [JMP (-(size cb + size cc + 1))])"
--- a/src/HOL/IMP/Def_Init.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Def_Init.thy Tue Aug 13 16:25:47 2013 +0200
@@ -7,9 +7,9 @@
inductive D :: "vname set \<Rightarrow> com \<Rightarrow> vname set \<Rightarrow> bool" where
Skip: "D A SKIP A" |
Assign: "vars a \<subseteq> A \<Longrightarrow> D A (x ::= a) (insert x A)" |
-Seq: "\<lbrakk> D A\<^isub>1 c\<^isub>1 A\<^isub>2; D A\<^isub>2 c\<^isub>2 A\<^isub>3 \<rbrakk> \<Longrightarrow> D A\<^isub>1 (c\<^isub>1;; c\<^isub>2) A\<^isub>3" |
-If: "\<lbrakk> vars b \<subseteq> A; D A c\<^isub>1 A\<^isub>1; D A c\<^isub>2 A\<^isub>2 \<rbrakk> \<Longrightarrow>
- D A (IF b THEN c\<^isub>1 ELSE c\<^isub>2) (A\<^isub>1 Int A\<^isub>2)" |
+Seq: "\<lbrakk> D A\<^sub>1 c\<^sub>1 A\<^sub>2; D A\<^sub>2 c\<^sub>2 A\<^sub>3 \<rbrakk> \<Longrightarrow> D A\<^sub>1 (c\<^sub>1;; c\<^sub>2) A\<^sub>3" |
+If: "\<lbrakk> vars b \<subseteq> A; D A c\<^sub>1 A\<^sub>1; D A c\<^sub>2 A\<^sub>2 \<rbrakk> \<Longrightarrow>
+ D A (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (A\<^sub>1 Int A\<^sub>2)" |
While: "\<lbrakk> vars b \<subseteq> A; D A c A' \<rbrakk> \<Longrightarrow> D A (WHILE b DO c) A"
inductive_cases [elim!]:
--- a/src/HOL/IMP/Def_Init_Big.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Def_Init_Big.thy Tue Aug 13 16:25:47 2013 +0200
@@ -13,13 +13,13 @@
Skip: "(SKIP,s) \<Rightarrow> s" |
AssignNone: "aval a s = None \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> None" |
Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, Some s) \<Rightarrow> Some(s(x := Some i))" |
-Seq: "(c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2 \<Longrightarrow> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<Longrightarrow> (c\<^isub>1;;c\<^isub>2,s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+Seq: "(c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2 \<Longrightarrow> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s\<^sub>1) \<Rightarrow> s\<^sub>3" |
-IfNone: "bval b s = None \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> None" |
-IfTrue: "\<lbrakk> bval b s = Some True; (c\<^isub>1,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
- (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> s'" |
-IfFalse: "\<lbrakk> bval b s = Some False; (c\<^isub>2,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
- (IF b THEN c\<^isub>1 ELSE c\<^isub>2,Some s) \<Rightarrow> s'" |
+IfNone: "bval b s = None \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> None" |
+IfTrue: "\<lbrakk> bval b s = Some True; (c\<^sub>1,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
+ (IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> s'" |
+IfFalse: "\<lbrakk> bval b s = Some False; (c\<^sub>2,Some s) \<Rightarrow> s' \<rbrakk> \<Longrightarrow>
+ (IF b THEN c\<^sub>1 ELSE c\<^sub>2,Some s) \<Rightarrow> s'" |
WhileNone: "bval b s = None \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> None" |
WhileFalse: "bval b s = Some False \<Longrightarrow> (WHILE b DO c,Some s) \<Rightarrow> Some s" |
--- a/src/HOL/IMP/Def_Init_Exp.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Def_Init_Exp.thy Tue Aug 13 16:25:47 2013 +0200
@@ -12,18 +12,18 @@
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> val option" where
"aval (N i) s = Some i" |
"aval (V x) s = s x" |
-"aval (Plus a\<^isub>1 a\<^isub>2) s =
- (case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
- (Some i\<^isub>1,Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1+i\<^isub>2) | _ \<Rightarrow> None)"
+"aval (Plus a\<^sub>1 a\<^sub>2) s =
+ (case (aval a\<^sub>1 s, aval a\<^sub>2 s) of
+ (Some i\<^sub>1,Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1+i\<^sub>2) | _ \<Rightarrow> None)"
fun bval :: "bexp \<Rightarrow> state \<Rightarrow> bool option" where
"bval (Bc v) s = Some v" |
"bval (Not b) s = (case bval b s of None \<Rightarrow> None | Some bv \<Rightarrow> Some(\<not> bv))" |
-"bval (And b\<^isub>1 b\<^isub>2) s = (case (bval b\<^isub>1 s, bval b\<^isub>2 s) of
- (Some bv\<^isub>1, Some bv\<^isub>2) \<Rightarrow> Some(bv\<^isub>1 & bv\<^isub>2) | _ \<Rightarrow> None)" |
-"bval (Less a\<^isub>1 a\<^isub>2) s = (case (aval a\<^isub>1 s, aval a\<^isub>2 s) of
- (Some i\<^isub>1, Some i\<^isub>2) \<Rightarrow> Some(i\<^isub>1 < i\<^isub>2) | _ \<Rightarrow> None)"
+"bval (And b\<^sub>1 b\<^sub>2) s = (case (bval b\<^sub>1 s, bval b\<^sub>2 s) of
+ (Some bv\<^sub>1, Some bv\<^sub>2) \<Rightarrow> Some(bv\<^sub>1 & bv\<^sub>2) | _ \<Rightarrow> None)" |
+"bval (Less a\<^sub>1 a\<^sub>2) s = (case (aval a\<^sub>1 s, aval a\<^sub>2 s) of
+ (Some i\<^sub>1, Some i\<^sub>2) \<Rightarrow> Some(i\<^sub>1 < i\<^sub>2) | _ \<Rightarrow> None)"
lemma aval_Some: "vars a \<subseteq> dom s \<Longrightarrow> \<exists> i. aval a s = Some i"
--- a/src/HOL/IMP/Def_Init_Small.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Def_Init_Small.thy Tue Aug 13 16:25:47 2013 +0200
@@ -12,10 +12,10 @@
Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := Some i))" |
Seq1: "(SKIP;;c,s) \<rightarrow> (c,s)" |
-Seq2: "(c\<^isub>1,s) \<rightarrow> (c\<^isub>1',s') \<Longrightarrow> (c\<^isub>1;;c\<^isub>2,s) \<rightarrow> (c\<^isub>1';;c\<^isub>2,s')" |
+Seq2: "(c\<^sub>1,s) \<rightarrow> (c\<^sub>1',s') \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s) \<rightarrow> (c\<^sub>1';;c\<^sub>2,s')" |
-IfTrue: "bval b s = Some True \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \<rightarrow> (c\<^isub>1,s)" |
-IfFalse: "bval b s = Some False \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \<rightarrow> (c\<^isub>2,s)" |
+IfTrue: "bval b s = Some True \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>1,s)" |
+IfFalse: "bval b s = Some False \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" |
While: "(WHILE b DO c,s) \<rightarrow> (IF b THEN c;; WHILE b DO c ELSE SKIP,s)"
--- a/src/HOL/IMP/Hoare.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Hoare.thy Tue Aug 13 16:25:47 2013 +0200
@@ -23,11 +23,11 @@
Assign: "\<turnstile> {\<lambda>s. P(s[a/x])} x::=a {P}" |
-Seq: "\<lbrakk> \<turnstile> {P} c\<^isub>1 {Q}; \<turnstile> {Q} c\<^isub>2 {R} \<rbrakk>
- \<Longrightarrow> \<turnstile> {P} c\<^isub>1;;c\<^isub>2 {R}" |
+Seq: "\<lbrakk> \<turnstile> {P} c\<^sub>1 {Q}; \<turnstile> {Q} c\<^sub>2 {R} \<rbrakk>
+ \<Longrightarrow> \<turnstile> {P} c\<^sub>1;;c\<^sub>2 {R}" |
-If: "\<lbrakk> \<turnstile> {\<lambda>s. P s \<and> bval b s} c\<^isub>1 {Q}; \<turnstile> {\<lambda>s. P s \<and> \<not> bval b s} c\<^isub>2 {Q} \<rbrakk>
- \<Longrightarrow> \<turnstile> {P} IF b THEN c\<^isub>1 ELSE c\<^isub>2 {Q}" |
+If: "\<lbrakk> \<turnstile> {\<lambda>s. P s \<and> bval b s} c\<^sub>1 {Q}; \<turnstile> {\<lambda>s. P s \<and> \<not> bval b s} c\<^sub>2 {Q} \<rbrakk>
+ \<Longrightarrow> \<turnstile> {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}" |
While: "\<turnstile> {\<lambda>s. P s \<and> bval b s} c {P} \<Longrightarrow>
\<turnstile> {P} WHILE b DO c {\<lambda>s. P s \<and> \<not> bval b s}" |
--- a/src/HOL/IMP/Hoare_Sound_Complete.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Hoare_Sound_Complete.thy Tue Aug 13 16:25:47 2013 +0200
@@ -31,12 +31,12 @@
lemma wp_Ass[simp]: "wp (x::=a) Q = (\<lambda>s. Q(s[a/x]))"
by (rule ext) (auto simp: wp_def)
-lemma wp_Seq[simp]: "wp (c\<^isub>1;;c\<^isub>2) Q = wp c\<^isub>1 (wp c\<^isub>2 Q)"
+lemma wp_Seq[simp]: "wp (c\<^sub>1;;c\<^sub>2) Q = wp c\<^sub>1 (wp c\<^sub>2 Q)"
by (rule ext) (auto simp: wp_def)
lemma wp_If[simp]:
- "wp (IF b THEN c\<^isub>1 ELSE c\<^isub>2) Q =
- (\<lambda>s. if bval b s then wp c\<^isub>1 Q s else wp c\<^isub>2 Q s)"
+ "wp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q =
+ (\<lambda>s. if bval b s then wp c\<^sub>1 Q s else wp c\<^sub>2 Q s)"
by (rule ext) (auto simp: wp_def)
lemma wp_While_If:
--- a/src/HOL/IMP/Hoare_Total.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Hoare_Total.thy Tue Aug 13 16:25:47 2013 +0200
@@ -25,10 +25,10 @@
Assign: "\<turnstile>\<^sub>t {\<lambda>s. P(s[a/x])} x::=a {P}" |
-Seq: "\<lbrakk> \<turnstile>\<^sub>t {P\<^isub>1} c\<^isub>1 {P\<^isub>2}; \<turnstile>\<^sub>t {P\<^isub>2} c\<^isub>2 {P\<^isub>3} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P\<^isub>1} c\<^isub>1;;c\<^isub>2 {P\<^isub>3}" |
+Seq: "\<lbrakk> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1 {P\<^sub>2}; \<turnstile>\<^sub>t {P\<^sub>2} c\<^sub>2 {P\<^sub>3} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}" |
-If: "\<lbrakk> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s} c\<^isub>1 {Q}; \<turnstile>\<^sub>t {\<lambda>s. P s \<and> \<not> bval b s} c\<^isub>2 {Q} \<rbrakk>
- \<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^isub>1 ELSE c\<^isub>2 {Q}" |
+If: "\<lbrakk> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s} c\<^sub>1 {Q}; \<turnstile>\<^sub>t {\<lambda>s. P s \<and> \<not> bval b s} c\<^sub>2 {Q} \<rbrakk>
+ \<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}" |
While:
"(\<And>n::nat.
@@ -119,14 +119,14 @@
lemma [simp]: "wp\<^sub>t (x ::= e) Q = (\<lambda>s. Q(s(x := aval e s)))"
by(auto intro!: ext simp: wpt_def)
-lemma [simp]: "wp\<^sub>t (c\<^isub>1;;c\<^isub>2) Q = wp\<^sub>t c\<^isub>1 (wp\<^sub>t c\<^isub>2 Q)"
+lemma [simp]: "wp\<^sub>t (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>t c\<^sub>1 (wp\<^sub>t c\<^sub>2 Q)"
unfolding wpt_def
apply(rule ext)
apply auto
done
lemma [simp]:
- "wp\<^sub>t (IF b THEN c\<^isub>1 ELSE c\<^isub>2) Q = (\<lambda>s. wp\<^sub>t (if bval b s then c\<^isub>1 else c\<^isub>2) Q s)"
+ "wp\<^sub>t (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\<lambda>s. wp\<^sub>t (if bval b s then c\<^sub>1 else c\<^sub>2) Q s)"
apply(unfold wpt_def)
apply(rule ext)
apply auto
--- a/src/HOL/IMP/Live.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Live.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,8 +10,8 @@
fun L :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
"L SKIP X = X" |
"L (x ::= a) X = vars a \<union> (X - {x})" |
-"L (c\<^isub>1;; c\<^isub>2) X = L c\<^isub>1 (L c\<^isub>2 X)" |
-"L (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = vars b \<union> L c\<^isub>1 X \<union> L c\<^isub>2 X" |
+"L (c\<^sub>1;; c\<^sub>2) X = L c\<^sub>1 (L c\<^sub>2 X)" |
+"L (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = vars b \<union> L c\<^sub>1 X \<union> L c\<^sub>2 X" |
"L (WHILE b DO c) X = vars b \<union> X \<union> L c X"
value "show (L (''y'' ::= V ''z'';; ''x'' ::= Plus (V ''y'') (V ''z'')) {''x''})"
@@ -21,15 +21,15 @@
fun "kill" :: "com \<Rightarrow> vname set" where
"kill SKIP = {}" |
"kill (x ::= a) = {x}" |
-"kill (c\<^isub>1;; c\<^isub>2) = kill c\<^isub>1 \<union> kill c\<^isub>2" |
-"kill (IF b THEN c\<^isub>1 ELSE c\<^isub>2) = kill c\<^isub>1 \<inter> kill c\<^isub>2" |
+"kill (c\<^sub>1;; c\<^sub>2) = kill c\<^sub>1 \<union> kill c\<^sub>2" |
+"kill (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = kill c\<^sub>1 \<inter> kill c\<^sub>2" |
"kill (WHILE b DO c) = {}"
fun gen :: "com \<Rightarrow> vname set" where
"gen SKIP = {}" |
"gen (x ::= a) = vars a" |
-"gen (c\<^isub>1;; c\<^isub>2) = gen c\<^isub>1 \<union> (gen c\<^isub>2 - kill c\<^isub>1)" |
-"gen (IF b THEN c\<^isub>1 ELSE c\<^isub>2) = vars b \<union> gen c\<^isub>1 \<union> gen c\<^isub>2" |
+"gen (c\<^sub>1;; c\<^sub>2) = gen c\<^sub>1 \<union> (gen c\<^sub>2 - kill c\<^sub>1)" |
+"gen (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = vars b \<union> gen c\<^sub>1 \<union> gen c\<^sub>2" |
"gen (WHILE b DO c) = vars b \<union> gen c"
lemma L_gen_kill: "L c X = gen c \<union> (X - kill c)"
@@ -110,8 +110,8 @@
fun bury :: "com \<Rightarrow> vname set \<Rightarrow> com" where
"bury SKIP X = SKIP" |
"bury (x ::= a) X = (if x \<in> X then x ::= a else SKIP)" |
-"bury (c\<^isub>1;; c\<^isub>2) X = (bury c\<^isub>1 (L c\<^isub>2 X);; bury c\<^isub>2 X)" |
-"bury (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = IF b THEN bury c\<^isub>1 X ELSE bury c\<^isub>2 X" |
+"bury (c\<^sub>1;; c\<^sub>2) X = (bury c\<^sub>1 (L c\<^sub>2 X);; bury c\<^sub>2 X)" |
+"bury (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = IF b THEN bury c\<^sub>1 X ELSE bury c\<^sub>2 X" |
"bury (WHILE b DO c) X = WHILE b DO bury c (L (WHILE b DO c) X)"
text{* We could prove the analogous lemma to @{thm[source]L_correct}, and the
@@ -182,8 +182,8 @@
lemma Assign_bury[simp]: "x::=a = bury c X \<longleftrightarrow> c = x::=a & x : X"
by (cases c) auto
-lemma Seq_bury[simp]: "bc\<^isub>1;;bc\<^isub>2 = bury c X \<longleftrightarrow>
- (EX c\<^isub>1 c\<^isub>2. c = c\<^isub>1;;c\<^isub>2 & bc\<^isub>2 = bury c\<^isub>2 X & bc\<^isub>1 = bury c\<^isub>1 (L c\<^isub>2 X))"
+lemma Seq_bury[simp]: "bc\<^sub>1;;bc\<^sub>2 = bury c X \<longleftrightarrow>
+ (EX c\<^sub>1 c\<^sub>2. c = c\<^sub>1;;c\<^sub>2 & bc\<^sub>2 = bury c\<^sub>2 X & bc\<^sub>1 = bury c\<^sub>1 (L c\<^sub>2 X))"
by (cases c) auto
lemma If_bury[simp]: "IF b THEN bc1 ELSE bc2 = bury c X \<longleftrightarrow>
--- a/src/HOL/IMP/Live_True.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Live_True.thy Tue Aug 13 16:25:47 2013 +0200
@@ -9,8 +9,8 @@
fun L :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
"L SKIP X = X" |
"L (x ::= a) X = (if x \<in> X then vars a \<union> (X - {x}) else X)" |
-"L (c\<^isub>1;; c\<^isub>2) X = L c\<^isub>1 (L c\<^isub>2 X)" |
-"L (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = vars b \<union> L c\<^isub>1 X \<union> L c\<^isub>2 X" |
+"L (c\<^sub>1;; c\<^sub>2) X = L c\<^sub>1 (L c\<^sub>2 X)" |
+"L (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = vars b \<union> L c\<^sub>1 X \<union> L c\<^sub>2 X" |
"L (WHILE b DO c) X = lfp(\<lambda>Y. vars b \<union> X \<union> L c Y)"
lemma L_mono: "mono (L c)"
@@ -171,8 +171,8 @@
fun Lb :: "com \<Rightarrow> vname set \<Rightarrow> vname set" where
"Lb SKIP X = X" |
"Lb (x ::= a) X = (if x \<in> X then X - {x} \<union> vars a else X)" |
-"Lb (c\<^isub>1;; c\<^isub>2) X = (Lb c\<^isub>1 \<circ> Lb c\<^isub>2) X" |
-"Lb (IF b THEN c\<^isub>1 ELSE c\<^isub>2) X = vars b \<union> Lb c\<^isub>1 X \<union> Lb c\<^isub>2 X" |
+"Lb (c\<^sub>1;; c\<^sub>2) X = (Lb c\<^sub>1 \<circ> Lb c\<^sub>2) X" |
+"Lb (IF b THEN c\<^sub>1 ELSE c\<^sub>2) X = vars b \<union> Lb c\<^sub>1 X \<union> Lb c\<^sub>2 X" |
"Lb (WHILE b DO c) X = iter (\<lambda>A. vars b \<union> X \<union> Lb c A) 2 {} (vars b \<union> rvars c \<union> X)"
text{* @{const Lb} (and @{const iter}) is not monotone! *}
--- a/src/HOL/IMP/OO.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/OO.thy Tue Aug 13 16:25:47 2013 +0200
@@ -48,33 +48,33 @@
"me \<turnstile> (e,c) \<Rightarrow> (r,ve',sn') \<Longrightarrow>
me \<turnstile> (x ::= e,c) \<Rightarrow> (r,ve'(x:=r),sn')" |
Fassign:
-"\<lbrakk> me \<turnstile> (oe,c\<^isub>1) \<Rightarrow> (Ref a,c\<^isub>2); me \<turnstile> (e,c\<^isub>2) \<Rightarrow> (r,ve\<^isub>3,s\<^isub>3,n\<^isub>3) \<rbrakk> \<Longrightarrow>
- me \<turnstile> (oe\<bullet>f ::= e,c\<^isub>1) \<Rightarrow> (r,ve\<^isub>3,s\<^isub>3(a,f := r),n\<^isub>3)" |
+"\<lbrakk> me \<turnstile> (oe,c\<^sub>1) \<Rightarrow> (Ref a,c\<^sub>2); me \<turnstile> (e,c\<^sub>2) \<Rightarrow> (r,ve\<^sub>3,s\<^sub>3,n\<^sub>3) \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (oe\<bullet>f ::= e,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,s\<^sub>3(a,f := r),n\<^sub>3)" |
Mcall:
-"\<lbrakk> me \<turnstile> (oe,c\<^isub>1) \<Rightarrow> (or,c\<^isub>2); me \<turnstile> (pe,c\<^isub>2) \<Rightarrow> (pr,ve\<^isub>3,sn\<^isub>3);
+"\<lbrakk> me \<turnstile> (oe,c\<^sub>1) \<Rightarrow> (or,c\<^sub>2); me \<turnstile> (pe,c\<^sub>2) \<Rightarrow> (pr,ve\<^sub>3,sn\<^sub>3);
ve = (\<lambda>x. null)(''this'' := or, ''param'' := pr);
- me \<turnstile> (me m,ve,sn\<^isub>3) \<Rightarrow> (r,ve',sn\<^isub>4) \<rbrakk>
+ me \<turnstile> (me m,ve,sn\<^sub>3) \<Rightarrow> (r,ve',sn\<^sub>4) \<rbrakk>
\<Longrightarrow>
- me \<turnstile> (oe\<bullet>m<pe>,c\<^isub>1) \<Rightarrow> (r,ve\<^isub>3,sn\<^isub>4)" |
+ me \<turnstile> (oe\<bullet>m<pe>,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,sn\<^sub>4)" |
Seq:
-"\<lbrakk> me \<turnstile> (e\<^isub>1,c\<^isub>1) \<Rightarrow> (r,c\<^isub>2); me \<turnstile> (e\<^isub>2,c\<^isub>2) \<Rightarrow> c\<^isub>3 \<rbrakk> \<Longrightarrow>
- me \<turnstile> (e\<^isub>1; e\<^isub>2,c\<^isub>1) \<Rightarrow> c\<^isub>3" |
+"\<lbrakk> me \<turnstile> (e\<^sub>1,c\<^sub>1) \<Rightarrow> (r,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (e\<^sub>1; e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |
IfTrue:
-"\<lbrakk> me \<turnstile> (b,c\<^isub>1) \<rightarrow> (True,c\<^isub>2); me \<turnstile> (e\<^isub>1,c\<^isub>2) \<Rightarrow> c\<^isub>3 \<rbrakk> \<Longrightarrow>
- me \<turnstile> (IF b THEN e\<^isub>1 ELSE e\<^isub>2,c\<^isub>1) \<Rightarrow> c\<^isub>3" |
+"\<lbrakk> me \<turnstile> (b,c\<^sub>1) \<rightarrow> (True,c\<^sub>2); me \<turnstile> (e\<^sub>1,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (IF b THEN e\<^sub>1 ELSE e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |
IfFalse:
-"\<lbrakk> me \<turnstile> (b,c\<^isub>1) \<rightarrow> (False,c\<^isub>2); me \<turnstile> (e\<^isub>2,c\<^isub>2) \<Rightarrow> c\<^isub>3 \<rbrakk> \<Longrightarrow>
- me \<turnstile> (IF b THEN e\<^isub>1 ELSE e\<^isub>2,c\<^isub>1) \<Rightarrow> c\<^isub>3" |
+"\<lbrakk> me \<turnstile> (b,c\<^sub>1) \<rightarrow> (False,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (IF b THEN e\<^sub>1 ELSE e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |
"me \<turnstile> (B bv,c) \<rightarrow> (bv,c)" |
-"me \<turnstile> (b,c\<^isub>1) \<rightarrow> (bv,c\<^isub>2) \<Longrightarrow> me \<turnstile> (Not b,c\<^isub>1) \<rightarrow> (\<not>bv,c\<^isub>2)" |
+"me \<turnstile> (b,c\<^sub>1) \<rightarrow> (bv,c\<^sub>2) \<Longrightarrow> me \<turnstile> (Not b,c\<^sub>1) \<rightarrow> (\<not>bv,c\<^sub>2)" |
-"\<lbrakk> me \<turnstile> (b\<^isub>1,c\<^isub>1) \<rightarrow> (bv\<^isub>1,c\<^isub>2); me \<turnstile> (b\<^isub>2,c\<^isub>2) \<rightarrow> (bv\<^isub>2,c\<^isub>3) \<rbrakk> \<Longrightarrow>
- me \<turnstile> (And b\<^isub>1 b\<^isub>2,c\<^isub>1) \<rightarrow> (bv\<^isub>1\<and>bv\<^isub>2,c\<^isub>3)" |
+"\<lbrakk> me \<turnstile> (b\<^sub>1,c\<^sub>1) \<rightarrow> (bv\<^sub>1,c\<^sub>2); me \<turnstile> (b\<^sub>2,c\<^sub>2) \<rightarrow> (bv\<^sub>2,c\<^sub>3) \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (And b\<^sub>1 b\<^sub>2,c\<^sub>1) \<rightarrow> (bv\<^sub>1\<and>bv\<^sub>2,c\<^sub>3)" |
-"\<lbrakk> me \<turnstile> (e\<^isub>1,c\<^isub>1) \<Rightarrow> (r\<^isub>1,c\<^isub>2); me \<turnstile> (e\<^isub>2,c\<^isub>2) \<Rightarrow> (r\<^isub>2,c\<^isub>3) \<rbrakk> \<Longrightarrow>
- me \<turnstile> (Eq e\<^isub>1 e\<^isub>2,c\<^isub>1) \<rightarrow> (r\<^isub>1=r\<^isub>2,c\<^isub>3)"
+"\<lbrakk> me \<turnstile> (e\<^sub>1,c\<^sub>1) \<Rightarrow> (r\<^sub>1,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> (r\<^sub>2,c\<^sub>3) \<rbrakk> \<Longrightarrow>
+ me \<turnstile> (Eq e\<^sub>1 e\<^sub>2,c\<^sub>1) \<rightarrow> (r\<^sub>1=r\<^sub>2,c\<^sub>3)"
code_pred (modes: i => i => o => bool) big_step .
--- a/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,18 +10,18 @@
where
Skip: "pe \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "pe \<turnstile> (x ::= a,s) \<Rightarrow> s(x := aval a s)" |
-Seq: "\<lbrakk> pe \<turnstile> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (c\<^isub>1;;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+Seq: "\<lbrakk> pe \<turnstile> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
-IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
-IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
+IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^sub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
+IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^sub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
WhileFalse: "\<not>bval b s \<Longrightarrow> pe \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" |
WhileTrue:
- "\<lbrakk> bval b s\<^isub>1; pe \<turnstile> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+ "\<lbrakk> bval b s\<^sub>1; pe \<turnstile> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
Var: "pe \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow> pe \<turnstile> ({VAR x; c}, s) \<Rightarrow> t(x := s x)" |
--- a/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,18 +10,18 @@
where
Skip: "pe \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "pe \<turnstile> (x ::= a,s) \<Rightarrow> s(x := aval a s)" |
-Seq: "\<lbrakk> pe \<turnstile> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (c\<^isub>1;;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+Seq: "\<lbrakk> pe \<turnstile> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
-IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
-IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
+IfTrue: "\<lbrakk> bval b s; pe \<turnstile> (c\<^sub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
+IfFalse: "\<lbrakk> \<not>bval b s; pe \<turnstile> (c\<^sub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
WhileFalse: "\<not>bval b s \<Longrightarrow> pe \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" |
WhileTrue:
- "\<lbrakk> bval b s\<^isub>1; pe \<turnstile> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2; pe \<turnstile> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- pe \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+ "\<lbrakk> bval b s\<^sub>1; pe \<turnstile> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2; pe \<turnstile> (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ pe \<turnstile> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
Var: "pe \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow> pe \<turnstile> ({VAR x; c}, s) \<Rightarrow> t(x := s x)" |
--- a/src/HOL/IMP/Procs_Stat_Vars_Stat.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Procs_Stat_Vars_Stat.thy Tue Aug 13 16:25:47 2013 +0200
@@ -17,19 +17,19 @@
where
Skip: "e \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "(pe,ve,f) \<turnstile> (x ::= a,s) \<Rightarrow> s(ve x := aval a (s o ve))" |
-Seq: "\<lbrakk> e \<turnstile> (c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2; e \<turnstile> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- e \<turnstile> (c\<^isub>1;;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+Seq: "\<lbrakk> e \<turnstile> (c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2; e \<turnstile> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ e \<turnstile> (c\<^sub>1;;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
-IfTrue: "\<lbrakk> bval b (s \<circ> venv e); e \<turnstile> (c\<^isub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- e \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
-IfFalse: "\<lbrakk> \<not>bval b (s \<circ> venv e); e \<turnstile> (c\<^isub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
- e \<turnstile> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t" |
+IfTrue: "\<lbrakk> bval b (s \<circ> venv e); e \<turnstile> (c\<^sub>1,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ e \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
+IfFalse: "\<lbrakk> \<not>bval b (s \<circ> venv e); e \<turnstile> (c\<^sub>2,s) \<Rightarrow> t \<rbrakk> \<Longrightarrow>
+ e \<turnstile> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t" |
WhileFalse: "\<not>bval b (s \<circ> venv e) \<Longrightarrow> e \<turnstile> (WHILE b DO c,s) \<Rightarrow> s" |
WhileTrue:
- "\<lbrakk> bval b (s\<^isub>1 \<circ> venv e); e \<turnstile> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2;
- e \<turnstile> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3 \<rbrakk> \<Longrightarrow>
- e \<turnstile> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3" |
+ "\<lbrakk> bval b (s\<^sub>1 \<circ> venv e); e \<turnstile> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2;
+ e \<turnstile> (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3 \<rbrakk> \<Longrightarrow>
+ e \<turnstile> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3" |
Var: "(pe,ve(x:=f),f+1) \<turnstile> (c,s) \<Rightarrow> t \<Longrightarrow>
(pe,ve,f) \<turnstile> ({VAR x; c}, s) \<Rightarrow> t" |
--- a/src/HOL/IMP/Sec_Type_Expr.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Sec_Type_Expr.thy Tue Aug 13 16:25:47 2013 +0200
@@ -29,7 +29,7 @@
fun sec_aexp :: "aexp \<Rightarrow> level" where
"sec (N n) = 0" |
"sec (V x) = sec x" |
-"sec (Plus a\<^isub>1 a\<^isub>2) = max (sec a\<^isub>1) (sec a\<^isub>2)"
+"sec (Plus a\<^sub>1 a\<^sub>2) = max (sec a\<^sub>1) (sec a\<^sub>2)"
instance ..
@@ -41,8 +41,8 @@
fun sec_bexp :: "bexp \<Rightarrow> level" where
"sec (Bc v) = 0" |
"sec (Not b) = sec b" |
-"sec (And b\<^isub>1 b\<^isub>2) = max (sec b\<^isub>1) (sec b\<^isub>2)" |
-"sec (Less a\<^isub>1 a\<^isub>2) = max (sec a\<^isub>1) (sec a\<^isub>2)"
+"sec (And b\<^sub>1 b\<^sub>2) = max (sec b\<^sub>1) (sec b\<^sub>2)" |
+"sec (Less a\<^sub>1 a\<^sub>2) = max (sec a\<^sub>1) (sec a\<^sub>2)"
instance ..
@@ -58,11 +58,11 @@
"s = s' (< l) == (\<forall> x. sec x < l \<longrightarrow> s x = s' x)"
lemma aval_eq_if_eq_le:
- "\<lbrakk> s\<^isub>1 = s\<^isub>2 (\<le> l); sec a \<le> l \<rbrakk> \<Longrightarrow> aval a s\<^isub>1 = aval a s\<^isub>2"
+ "\<lbrakk> s\<^sub>1 = s\<^sub>2 (\<le> l); sec a \<le> l \<rbrakk> \<Longrightarrow> aval a s\<^sub>1 = aval a s\<^sub>2"
by (induct a) auto
lemma bval_eq_if_eq_le:
- "\<lbrakk> s\<^isub>1 = s\<^isub>2 (\<le> l); sec b \<le> l \<rbrakk> \<Longrightarrow> bval b s\<^isub>1 = bval b s\<^isub>2"
+ "\<lbrakk> s\<^sub>1 = s\<^sub>2 (\<le> l); sec b \<le> l \<rbrakk> \<Longrightarrow> bval b s\<^sub>1 = bval b s\<^sub>2"
by (induct b) (auto simp add: aval_eq_if_eq_le)
end
--- a/src/HOL/IMP/Sec_Typing.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Sec_Typing.thy Tue Aug 13 16:25:47 2013 +0200
@@ -11,9 +11,9 @@
Assign:
"\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile> x ::= a" |
Seq:
- "\<lbrakk> l \<turnstile> c\<^isub>1; l \<turnstile> c\<^isub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> c\<^isub>1;;c\<^isub>2" |
+ "\<lbrakk> l \<turnstile> c\<^sub>1; l \<turnstile> c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> c\<^sub>1;;c\<^sub>2" |
If:
- "\<lbrakk> max (sec b) l \<turnstile> c\<^isub>1; max (sec b) l \<turnstile> c\<^isub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> IF b THEN c\<^isub>1 ELSE c\<^isub>2" |
+ "\<lbrakk> max (sec b) l \<turnstile> c\<^sub>1; max (sec b) l \<turnstile> c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" |
While:
"max (sec b) l \<turnstile> c \<Longrightarrow> l \<turnstile> WHILE b DO c"
@@ -24,7 +24,7 @@
value "2 \<turnstile> IF Less (V ''x1'') (V ''x'') THEN ''x1'' ::= N 0 ELSE SKIP"
inductive_cases [elim!]:
- "l \<turnstile> x ::= a" "l \<turnstile> c\<^isub>1;;c\<^isub>2" "l \<turnstile> IF b THEN c\<^isub>1 ELSE c\<^isub>2" "l \<turnstile> WHILE b DO c"
+ "l \<turnstile> x ::= a" "l \<turnstile> c\<^sub>1;;c\<^sub>2" "l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" "l \<turnstile> WHILE b DO c"
text{* An important property: anti-monotonicity. *}
@@ -187,9 +187,9 @@
Assign':
"\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile>' x ::= a" |
Seq':
- "\<lbrakk> l \<turnstile>' c\<^isub>1; l \<turnstile>' c\<^isub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' c\<^isub>1;;c\<^isub>2" |
+ "\<lbrakk> l \<turnstile>' c\<^sub>1; l \<turnstile>' c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' c\<^sub>1;;c\<^sub>2" |
If':
- "\<lbrakk> sec b \<le> l; l \<turnstile>' c\<^isub>1; l \<turnstile>' c\<^isub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' IF b THEN c\<^isub>1 ELSE c\<^isub>2" |
+ "\<lbrakk> sec b \<le> l; l \<turnstile>' c\<^sub>1; l \<turnstile>' c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' IF b THEN c\<^sub>1 ELSE c\<^sub>2" |
While':
"\<lbrakk> sec b \<le> l; l \<turnstile>' c \<rbrakk> \<Longrightarrow> l \<turnstile>' WHILE b DO c" |
anti_mono':
@@ -221,10 +221,10 @@
Assign2:
"sec x \<ge> sec a \<Longrightarrow> \<turnstile> x ::= a : sec x" |
Seq2:
- "\<lbrakk> \<turnstile> c\<^isub>1 : l\<^isub>1; \<turnstile> c\<^isub>2 : l\<^isub>2 \<rbrakk> \<Longrightarrow> \<turnstile> c\<^isub>1;;c\<^isub>2 : min l\<^isub>1 l\<^isub>2 " |
+ "\<lbrakk> \<turnstile> c\<^sub>1 : l\<^sub>1; \<turnstile> c\<^sub>2 : l\<^sub>2 \<rbrakk> \<Longrightarrow> \<turnstile> c\<^sub>1;;c\<^sub>2 : min l\<^sub>1 l\<^sub>2 " |
If2:
- "\<lbrakk> sec b \<le> min l\<^isub>1 l\<^isub>2; \<turnstile> c\<^isub>1 : l\<^isub>1; \<turnstile> c\<^isub>2 : l\<^isub>2 \<rbrakk>
- \<Longrightarrow> \<turnstile> IF b THEN c\<^isub>1 ELSE c\<^isub>2 : min l\<^isub>1 l\<^isub>2" |
+ "\<lbrakk> sec b \<le> min l\<^sub>1 l\<^sub>2; \<turnstile> c\<^sub>1 : l\<^sub>1; \<turnstile> c\<^sub>2 : l\<^sub>2 \<rbrakk>
+ \<Longrightarrow> \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2 : min l\<^sub>1 l\<^sub>2" |
While2:
"\<lbrakk> sec b \<le> l; \<turnstile> c : l \<rbrakk> \<Longrightarrow> \<turnstile> WHILE b DO c : l"
--- a/src/HOL/IMP/Sec_TypingT.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Sec_TypingT.thy Tue Aug 13 16:25:47 2013 +0200
@@ -9,17 +9,17 @@
Assign:
"\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile> x ::= a" |
Seq:
- "l \<turnstile> c\<^isub>1 \<Longrightarrow> l \<turnstile> c\<^isub>2 \<Longrightarrow> l \<turnstile> c\<^isub>1;;c\<^isub>2" |
+ "l \<turnstile> c\<^sub>1 \<Longrightarrow> l \<turnstile> c\<^sub>2 \<Longrightarrow> l \<turnstile> c\<^sub>1;;c\<^sub>2" |
If:
- "\<lbrakk> max (sec b) l \<turnstile> c\<^isub>1; max (sec b) l \<turnstile> c\<^isub>2 \<rbrakk>
- \<Longrightarrow> l \<turnstile> IF b THEN c\<^isub>1 ELSE c\<^isub>2" |
+ "\<lbrakk> max (sec b) l \<turnstile> c\<^sub>1; max (sec b) l \<turnstile> c\<^sub>2 \<rbrakk>
+ \<Longrightarrow> l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" |
While:
"sec b = 0 \<Longrightarrow> 0 \<turnstile> c \<Longrightarrow> 0 \<turnstile> WHILE b DO c"
code_pred (expected_modes: i => i => bool) sec_type .
inductive_cases [elim!]:
- "l \<turnstile> x ::= a" "l \<turnstile> c\<^isub>1;;c\<^isub>2" "l \<turnstile> IF b THEN c\<^isub>1 ELSE c\<^isub>2" "l \<turnstile> WHILE b DO c"
+ "l \<turnstile> x ::= a" "l \<turnstile> c\<^sub>1;;c\<^sub>2" "l \<turnstile> IF b THEN c\<^sub>1 ELSE c\<^sub>2" "l \<turnstile> WHILE b DO c"
lemma anti_mono: "l \<turnstile> c \<Longrightarrow> l' \<le> l \<Longrightarrow> l' \<turnstile> c"
@@ -176,9 +176,9 @@
Assign':
"\<lbrakk> sec x \<ge> sec a; sec x \<ge> l \<rbrakk> \<Longrightarrow> l \<turnstile>' x ::= a" |
Seq':
- "l \<turnstile>' c\<^isub>1 \<Longrightarrow> l \<turnstile>' c\<^isub>2 \<Longrightarrow> l \<turnstile>' c\<^isub>1;;c\<^isub>2" |
+ "l \<turnstile>' c\<^sub>1 \<Longrightarrow> l \<turnstile>' c\<^sub>2 \<Longrightarrow> l \<turnstile>' c\<^sub>1;;c\<^sub>2" |
If':
- "\<lbrakk> sec b \<le> l; l \<turnstile>' c\<^isub>1; l \<turnstile>' c\<^isub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' IF b THEN c\<^isub>1 ELSE c\<^isub>2" |
+ "\<lbrakk> sec b \<le> l; l \<turnstile>' c\<^sub>1; l \<turnstile>' c\<^sub>2 \<rbrakk> \<Longrightarrow> l \<turnstile>' IF b THEN c\<^sub>1 ELSE c\<^sub>2" |
While':
"\<lbrakk> sec b = 0; 0 \<turnstile>' c \<rbrakk> \<Longrightarrow> 0 \<turnstile>' WHILE b DO c" |
anti_mono':
--- a/src/HOL/IMP/Small_Step.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Small_Step.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,11 +10,11 @@
where
Assign: "(x ::= a, s) \<rightarrow> (SKIP, s(x := aval a s))" |
-Seq1: "(SKIP;;c\<^isub>2,s) \<rightarrow> (c\<^isub>2,s)" |
-Seq2: "(c\<^isub>1,s) \<rightarrow> (c\<^isub>1',s') \<Longrightarrow> (c\<^isub>1;;c\<^isub>2,s) \<rightarrow> (c\<^isub>1';;c\<^isub>2,s')" |
+Seq1: "(SKIP;;c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" |
+Seq2: "(c\<^sub>1,s) \<rightarrow> (c\<^sub>1',s') \<Longrightarrow> (c\<^sub>1;;c\<^sub>2,s) \<rightarrow> (c\<^sub>1';;c\<^sub>2,s')" |
-IfTrue: "bval b s \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \<rightarrow> (c\<^isub>1,s)" |
-IfFalse: "\<not>bval b s \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2,s) \<rightarrow> (c\<^isub>2,s)" |
+IfTrue: "bval b s \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>1,s)" |
+IfFalse: "\<not>bval b s \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2,s) \<rightarrow> (c\<^sub>2,s)" |
While: "(WHILE b DO c,s) \<rightarrow>
(IF b THEN c;; WHILE b DO c ELSE SKIP,s)"
--- a/src/HOL/IMP/VCG.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/VCG.thy Tue Aug 13 16:25:47 2013 +0200
@@ -20,8 +20,8 @@
fun strip :: "acom \<Rightarrow> com" where
"strip SKIP = com.SKIP" |
"strip (x ::= a) = (x ::= a)" |
-"strip (C\<^isub>1;; C\<^isub>2) = (strip C\<^isub>1;; strip C\<^isub>2)" |
-"strip (IF b THEN C\<^isub>1 ELSE C\<^isub>2) = (IF b THEN strip C\<^isub>1 ELSE strip C\<^isub>2)" |
+"strip (C\<^sub>1;; C\<^sub>2) = (strip C\<^sub>1;; strip C\<^sub>2)" |
+"strip (IF b THEN C\<^sub>1 ELSE C\<^sub>2) = (IF b THEN strip C\<^sub>1 ELSE strip C\<^sub>2)" |
"strip ({_} WHILE b DO C) = (WHILE b DO strip C)"
text{* Weakest precondition from annotated commands: *}
@@ -29,9 +29,9 @@
fun pre :: "acom \<Rightarrow> assn \<Rightarrow> assn" where
"pre SKIP Q = Q" |
"pre (x ::= a) Q = (\<lambda>s. Q(s(x := aval a s)))" |
-"pre (C\<^isub>1;; C\<^isub>2) Q = pre C\<^isub>1 (pre C\<^isub>2 Q)" |
-"pre (IF b THEN C\<^isub>1 ELSE C\<^isub>2) Q =
- (\<lambda>s. if bval b s then pre C\<^isub>1 Q s else pre C\<^isub>2 Q s)" |
+"pre (C\<^sub>1;; C\<^sub>2) Q = pre C\<^sub>1 (pre C\<^sub>2 Q)" |
+"pre (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q =
+ (\<lambda>s. if bval b s then pre C\<^sub>1 Q s else pre C\<^sub>2 Q s)" |
"pre ({I} WHILE b DO C) Q = I"
text{* Verification condition: *}
@@ -39,8 +39,8 @@
fun vc :: "acom \<Rightarrow> assn \<Rightarrow> assn" where
"vc SKIP Q = (\<lambda>s. True)" |
"vc (x ::= a) Q = (\<lambda>s. True)" |
-"vc (C\<^isub>1;; C\<^isub>2) Q = (\<lambda>s. vc C\<^isub>1 (pre C\<^isub>2 Q) s \<and> vc C\<^isub>2 Q s)" |
-"vc (IF b THEN C\<^isub>1 ELSE C\<^isub>2) Q = (\<lambda>s. vc C\<^isub>1 Q s \<and> vc C\<^isub>2 Q s)" |
+"vc (C\<^sub>1;; C\<^sub>2) Q = (\<lambda>s. vc C\<^sub>1 (pre C\<^sub>2 Q) s \<and> vc C\<^sub>2 Q s)" |
+"vc (IF b THEN C\<^sub>1 ELSE C\<^sub>2) Q = (\<lambda>s. vc C\<^sub>1 Q s \<and> vc C\<^sub>2 Q s)" |
"vc ({I} WHILE b DO C) Q =
(\<lambda>s. (I s \<and> bval b s \<longrightarrow> pre C I s) \<and>
(I s \<and> \<not> bval b s \<longrightarrow> Q s) \<and>
--- a/src/HOL/IMP/Vars.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/IMP/Vars.thy Tue Aug 13 16:25:47 2013 +0200
@@ -25,7 +25,7 @@
fun vars_aexp :: "aexp \<Rightarrow> vname set" where
"vars (N n) = {}" |
"vars (V x) = {x}" |
-"vars (Plus a\<^isub>1 a\<^isub>2) = vars a\<^isub>1 \<union> vars a\<^isub>2"
+"vars (Plus a\<^sub>1 a\<^sub>2) = vars a\<^sub>1 \<union> vars a\<^sub>2"
instance ..
@@ -39,8 +39,8 @@
fun vars_bexp :: "bexp \<Rightarrow> vname set" where
"vars (Bc v) = {}" |
"vars (Not b) = vars b" |
-"vars (And b\<^isub>1 b\<^isub>2) = vars b\<^isub>1 \<union> vars b\<^isub>2" |
-"vars (Less a\<^isub>1 a\<^isub>2) = vars a\<^isub>1 \<union> vars a\<^isub>2"
+"vars (And b\<^sub>1 b\<^sub>2) = vars b\<^sub>1 \<union> vars b\<^sub>2" |
+"vars (Less a\<^sub>1 a\<^sub>2) = vars a\<^sub>1 \<union> vars a\<^sub>2"
instance ..
@@ -54,16 +54,16 @@
"f = g on X == \<forall> x \<in> X. f x = g x"
lemma aval_eq_if_eq_on_vars[simp]:
- "s\<^isub>1 = s\<^isub>2 on vars a \<Longrightarrow> aval a s\<^isub>1 = aval a s\<^isub>2"
+ "s\<^sub>1 = s\<^sub>2 on vars a \<Longrightarrow> aval a s\<^sub>1 = aval a s\<^sub>2"
apply(induction a)
apply simp_all
done
lemma bval_eq_if_eq_on_vars:
- "s\<^isub>1 = s\<^isub>2 on vars b \<Longrightarrow> bval b s\<^isub>1 = bval b s\<^isub>2"
+ "s\<^sub>1 = s\<^sub>2 on vars b \<Longrightarrow> bval b s\<^sub>1 = bval b s\<^sub>2"
proof(induction b)
case (Less a1 a2)
- hence "aval a1 s\<^isub>1 = aval a1 s\<^isub>2" and "aval a2 s\<^isub>1 = aval a2 s\<^isub>2" by simp_all
+ hence "aval a1 s\<^sub>1 = aval a1 s\<^sub>2" and "aval a2 s\<^sub>1 = aval a2 s\<^sub>2" by simp_all
thus ?case by simp
qed simp_all
--- a/src/HOL/Induct/Ordinals.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Induct/Ordinals.thy Tue Aug 13 16:25:47 2013 +0200
@@ -56,7 +56,7 @@
| "veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))"
definition "veb a = veblen a Zero"
-definition "\<epsilon>\<^isub>0 = veb Zero"
-definition "\<Gamma>\<^isub>0 = Limit (\<lambda>n. iter veb n Zero)"
+definition "\<epsilon>\<^sub>0 = veb Zero"
+definition "\<Gamma>\<^sub>0 = Limit (\<lambda>n. iter veb n Zero)"
end
--- a/src/HOL/Library/Cardinality.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Cardinality.thy Tue Aug 13 16:25:47 2013 +0200
@@ -345,20 +345,20 @@
instance by intro_classes (simp add: card_UNIV_set_def card_UNIV_set card_UNIV)
end
-lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^isub>1]"
+lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^sub>1]"
by(auto intro: finite_1.exhaust)
-lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^isub>1, finite_2.a\<^isub>2]"
+lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^sub>1, finite_2.a\<^sub>2]"
by(auto intro: finite_2.exhaust)
-lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^isub>1, finite_3.a\<^isub>2, finite_3.a\<^isub>3]"
+lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^sub>1, finite_3.a\<^sub>2, finite_3.a\<^sub>3]"
by(auto intro: finite_3.exhaust)
-lemma UNIV_finite_4: "UNIV = set [finite_4.a\<^isub>1, finite_4.a\<^isub>2, finite_4.a\<^isub>3, finite_4.a\<^isub>4]"
+lemma UNIV_finite_4: "UNIV = set [finite_4.a\<^sub>1, finite_4.a\<^sub>2, finite_4.a\<^sub>3, finite_4.a\<^sub>4]"
by(auto intro: finite_4.exhaust)
lemma UNIV_finite_5:
- "UNIV = set [finite_5.a\<^isub>1, finite_5.a\<^isub>2, finite_5.a\<^isub>3, finite_5.a\<^isub>4, finite_5.a\<^isub>5]"
+ "UNIV = set [finite_5.a\<^sub>1, finite_5.a\<^sub>2, finite_5.a\<^sub>3, finite_5.a\<^sub>4, finite_5.a\<^sub>5]"
by(auto intro: finite_5.exhaust)
instantiation Enum.finite_1 :: card_UNIV begin
--- a/src/HOL/Library/Discrete.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Discrete.thy Tue Aug 13 16:25:47 2013 +0200
@@ -76,32 +76,32 @@
definition sqrt :: "nat \<Rightarrow> nat"
where
- "sqrt n = Max {m. m\<twosuperior> \<le> n}"
+ "sqrt n = Max {m. m\<^sup>2 \<le> n}"
lemma sqrt_aux:
fixes n :: nat
- shows "finite {m. m\<twosuperior> \<le> n}" and "{m. m\<twosuperior> \<le> n} \<noteq> {}"
+ shows "finite {m. m\<^sup>2 \<le> n}" and "{m. m\<^sup>2 \<le> n} \<noteq> {}"
proof -
{ fix m
- assume "m\<twosuperior> \<le> n"
+ assume "m\<^sup>2 \<le> n"
then have "m \<le> n"
by (cases m) (simp_all add: power2_eq_square)
} note ** = this
- then have "{m. m\<twosuperior> \<le> n} \<subseteq> {m. m \<le> n}" by auto
- then show "finite {m. m\<twosuperior> \<le> n}" by (rule finite_subset) rule
- have "0\<twosuperior> \<le> n" by simp
- then show *: "{m. m\<twosuperior> \<le> n} \<noteq> {}" by blast
+ then have "{m. m\<^sup>2 \<le> n} \<subseteq> {m. m \<le> n}" by auto
+ then show "finite {m. m\<^sup>2 \<le> n}" by (rule finite_subset) rule
+ have "0\<^sup>2 \<le> n" by simp
+ then show *: "{m. m\<^sup>2 \<le> n} \<noteq> {}" by blast
qed
lemma [code]:
- "sqrt n = Max (Set.filter (\<lambda>m. m\<twosuperior> \<le> n) {0..n})"
+ "sqrt n = Max (Set.filter (\<lambda>m. m\<^sup>2 \<le> n) {0..n})"
proof -
- from power2_nat_le_imp_le [of _ n] have "{m. m \<le> n \<and> m\<twosuperior> \<le> n} = {m. m\<twosuperior> \<le> n}" by auto
+ from power2_nat_le_imp_le [of _ n] have "{m. m \<le> n \<and> m\<^sup>2 \<le> n} = {m. m\<^sup>2 \<le> n}" by auto
then show ?thesis by (simp add: sqrt_def Set.filter_def)
qed
lemma sqrt_inverse_power2 [simp]:
- "sqrt (n\<twosuperior>) = n"
+ "sqrt (n\<^sup>2) = n"
proof -
have "{m. m \<le> n} \<noteq> {}" by auto
then have "Max {m. m \<le> n} \<le> n" by auto
@@ -121,30 +121,30 @@
lemma sqrt_greater_zero_iff [simp]:
"sqrt n > 0 \<longleftrightarrow> n > 0"
proof -
- have *: "0 < Max {m. m\<twosuperior> \<le> n} \<longleftrightarrow> (\<exists>a\<in>{m. m\<twosuperior> \<le> n}. 0 < a)"
+ have *: "0 < Max {m. m\<^sup>2 \<le> n} \<longleftrightarrow> (\<exists>a\<in>{m. m\<^sup>2 \<le> n}. 0 < a)"
by (rule Max_gr_iff) (fact sqrt_aux)+
show ?thesis
proof
assume "0 < sqrt n"
- then have "0 < Max {m. m\<twosuperior> \<le> n}" by (simp add: sqrt_def)
+ then have "0 < Max {m. m\<^sup>2 \<le> n}" by (simp add: sqrt_def)
with * show "0 < n" by (auto dest: power2_nat_le_imp_le)
next
assume "0 < n"
- then have "1\<twosuperior> \<le> n \<and> 0 < (1::nat)" by simp
- then have "\<exists>q. q\<twosuperior> \<le> n \<and> 0 < q" ..
- with * have "0 < Max {m. m\<twosuperior> \<le> n}" by blast
+ then have "1\<^sup>2 \<le> n \<and> 0 < (1::nat)" by simp
+ then have "\<exists>q. q\<^sup>2 \<le> n \<and> 0 < q" ..
+ with * have "0 < Max {m. m\<^sup>2 \<le> n}" by blast
then show "0 < sqrt n" by (simp add: sqrt_def)
qed
qed
lemma sqrt_power2_le [simp]: (* FIXME tune proof *)
- "(sqrt n)\<twosuperior> \<le> n"
+ "(sqrt n)\<^sup>2 \<le> n"
proof (cases "n > 0")
case False then show ?thesis by (simp add: sqrt_def)
next
case True then have "sqrt n > 0" by simp
- then have "mono (times (Max {m. m\<twosuperior> \<le> n}))" by (auto intro: mono_times_nat simp add: sqrt_def)
- then have *: "Max {m. m\<twosuperior> \<le> n} * Max {m. m\<twosuperior> \<le> n} = Max (times (Max {m. m\<twosuperior> \<le> n}) ` {m. m\<twosuperior> \<le> n})"
+ then have "mono (times (Max {m. m\<^sup>2 \<le> n}))" by (auto intro: mono_times_nat simp add: sqrt_def)
+ then have *: "Max {m. m\<^sup>2 \<le> n} * Max {m. m\<^sup>2 \<le> n} = Max (times (Max {m. m\<^sup>2 \<le> n}) ` {m. m\<^sup>2 \<le> n})"
using sqrt_aux [of n] by (rule mono_Max_commute)
have "Max (op * (Max {m. m * m \<le> n}) ` {m. m * m \<le> n}) \<le> n"
apply (subst Max_le_iff)
--- a/src/HOL/Library/FuncSet.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/FuncSet.thy Tue Aug 13 16:25:47 2013 +0200
@@ -342,21 +342,21 @@
definition PiE :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> ('a \<Rightarrow> 'b) set" where
"PiE S T = Pi S T \<inter> extensional S"
-abbreviation "Pi\<^isub>E A B \<equiv> PiE A B"
+abbreviation "Pi\<^sub>E A B \<equiv> PiE A B"
syntax "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10)
-syntax (xsymbols) "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+syntax (xsymbols) "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^sub>E _\<in>_./ _)" 10)
-syntax (HTML output) "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10)
+syntax (HTML output) "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^sub>E _\<in>_./ _)" 10)
-translations "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
+translations "PIE x:A. B" == "CONST Pi\<^sub>E A (%x. B)"
-abbreviation extensional_funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "->\<^isub>E" 60) where
- "A ->\<^isub>E B \<equiv> (\<Pi>\<^isub>E i\<in>A. B)"
+abbreviation extensional_funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "->\<^sub>E" 60) where
+ "A ->\<^sub>E B \<equiv> (\<Pi>\<^sub>E i\<in>A. B)"
notation (xsymbols)
- extensional_funcset (infixr "\<rightarrow>\<^isub>E" 60)
+ extensional_funcset (infixr "\<rightarrow>\<^sub>E" 60)
lemma extensional_funcset_def: "extensional_funcset S T = (S -> T) \<inter> extensional S"
by (simp add: PiE_def)
@@ -368,16 +368,16 @@
unfolding PiE_def by auto
lemma PiE_eq_empty_iff:
- "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
+ "Pi\<^sub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
proof
- assume "Pi\<^isub>E I F = {}"
+ assume "Pi\<^sub>E I F = {}"
show "\<exists>i\<in>I. F i = {}"
proof (rule ccontr)
assume "\<not> ?thesis"
then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto
from choice[OF this] guess f ..
- then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def PiE_def)
- with `Pi\<^isub>E I F = {}` show False by auto
+ then have "f \<in> Pi\<^sub>E I F" by (auto simp: extensional_def PiE_def)
+ with `Pi\<^sub>E I F = {}` show False by auto
qed
qed (auto simp: PiE_def)
@@ -405,11 +405,11 @@
then show ?thesis using assms by (auto intro: PiE_fun_upd)
qed
-lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)"
+lemma PiE_Int: "(Pi\<^sub>E I A) \<inter> (Pi\<^sub>E I B) = Pi\<^sub>E I (\<lambda>x. A x \<inter> B x)"
by (auto simp: PiE_def)
lemma PiE_cong:
- "(\<And>i. i\<in>I \<Longrightarrow> A i = B i) \<Longrightarrow> Pi\<^isub>E I A = Pi\<^isub>E I B"
+ "(\<And>i. i\<in>I \<Longrightarrow> A i = B i) \<Longrightarrow> Pi\<^sub>E I A = Pi\<^sub>E I B"
unfolding PiE_def by (auto simp: Pi_cong)
lemma PiE_E [elim]:
@@ -433,22 +433,22 @@
lemma PiE_eq_subset:
assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
- assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I"
+ assumes eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" and "i \<in> I"
shows "F i \<subseteq> F' i"
proof
fix x assume "x \<in> F i"
with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto
from choice[OF this] guess f .. note f = this
- then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def PiE_def)
- then have "f \<in> Pi\<^isub>E I F'" using assms by simp
+ then have "f \<in> Pi\<^sub>E I F" by (auto simp: extensional_def PiE_def)
+ then have "f \<in> Pi\<^sub>E I F'" using assms by simp
then show "x \<in> F' i" using f `i \<in> I` by (auto simp: PiE_def)
qed
lemma PiE_eq_iff_not_empty:
assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
- shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
+ shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
proof (intro iffI ballI)
- fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I"
+ fix i assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" and i: "i \<in> I"
show "F i = F' i"
using PiE_eq_subset[of I F F', OF ne eq i]
using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i]
@@ -456,21 +456,21 @@
qed (auto simp: PiE_def)
lemma PiE_eq_iff:
- "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
+ "Pi\<^sub>E I F = Pi\<^sub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
proof (intro iffI disjCI)
- assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+ assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'"
assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto
with PiE_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto
next
assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
- then show "Pi\<^isub>E I F = Pi\<^isub>E I F'"
+ then show "Pi\<^sub>E I F = Pi\<^sub>E I F'"
using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def)
qed
lemma extensional_funcset_fun_upd_restricts_rangeI:
- "\<forall>y \<in> S. f x \<noteq> f y \<Longrightarrow> f : (insert x S) \<rightarrow>\<^isub>E T ==> f(x := undefined) : S \<rightarrow>\<^isub>E (T - {f x})"
+ "\<forall>y \<in> S. f x \<noteq> f y \<Longrightarrow> f : (insert x S) \<rightarrow>\<^sub>E T ==> f(x := undefined) : S \<rightarrow>\<^sub>E (T - {f x})"
unfolding extensional_funcset_def extensional_def
apply auto
apply (case_tac "x = xa")
@@ -478,21 +478,21 @@
done
lemma extensional_funcset_fun_upd_extends_rangeI:
- assumes "a \<in> T" "f \<in> S \<rightarrow>\<^isub>E (T - {a})"
- shows "f(x := a) \<in> (insert x S) \<rightarrow>\<^isub>E T"
+ assumes "a \<in> T" "f \<in> S \<rightarrow>\<^sub>E (T - {a})"
+ shows "f(x := a) \<in> (insert x S) \<rightarrow>\<^sub>E T"
using assms unfolding extensional_funcset_def extensional_def by auto
subsubsection {* Injective Extensional Function Spaces *}
lemma extensional_funcset_fun_upd_inj_onI:
- assumes "f \<in> S \<rightarrow>\<^isub>E (T - {a})" "inj_on f S"
+ assumes "f \<in> S \<rightarrow>\<^sub>E (T - {a})" "inj_on f S"
shows "inj_on (f(x := a)) S"
using assms unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI)
lemma extensional_funcset_extend_domain_inj_on_eq:
assumes "x \<notin> S"
- shows"{f. f \<in> (insert x S) \<rightarrow>\<^isub>E T \<and> inj_on f (insert x S)} =
- (%(y, g). g(x:=y)) ` {(y, g). y \<in> T \<and> g \<in> S \<rightarrow>\<^isub>E (T - {y}) \<and> inj_on g S}"
+ shows"{f. f \<in> (insert x S) \<rightarrow>\<^sub>E T \<and> inj_on f (insert x S)} =
+ (%(y, g). g(x:=y)) ` {(y, g). y \<in> T \<and> g \<in> S \<rightarrow>\<^sub>E (T - {y}) \<and> inj_on g S}"
proof -
from assms show ?thesis
apply (auto del: PiE_I PiE_E)
@@ -508,7 +508,7 @@
lemma extensional_funcset_extend_domain_inj_onI:
assumes "x \<notin> S"
- shows "inj_on (\<lambda>(y, g). g(x := y)) {(y, g). y \<in> T \<and> g \<in> S \<rightarrow>\<^isub>E (T - {y}) \<and> inj_on g S}"
+ shows "inj_on (\<lambda>(y, g). g(x := y)) {(y, g). y \<in> T \<and> g \<in> S \<rightarrow>\<^sub>E (T - {y}) \<and> inj_on g S}"
proof -
from assms show ?thesis
apply (auto intro!: inj_onI)
@@ -522,9 +522,9 @@
lemma finite_PiE: "finite S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> finite (T i)) \<Longrightarrow> finite (PIE i : S. T i)"
by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq)
-lemma inj_combinator: "x \<notin> S \<Longrightarrow> inj_on (\<lambda>(y, g). g(x := y)) (T x \<times> Pi\<^isub>E S T)"
+lemma inj_combinator: "x \<notin> S \<Longrightarrow> inj_on (\<lambda>(y, g). g(x := y)) (T x \<times> Pi\<^sub>E S T)"
proof (safe intro!: inj_onI ext)
- fix f y g z assume "x \<notin> S" and fg: "f \<in> Pi\<^isub>E S T" "g \<in> Pi\<^isub>E S T"
+ fix f y g z assume "x \<notin> S" and fg: "f \<in> Pi\<^sub>E S T" "g \<in> Pi\<^sub>E S T"
assume "f(x := y) = g(x := z)"
then have *: "\<And>i. (f(x := y)) i = (g(x := z)) i"
unfolding fun_eq_iff by auto
--- a/src/HOL/Library/Function_Growth.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Function_Growth.thy Tue Aug 13 16:25:47 2013 +0200
@@ -73,35 +73,35 @@
definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<cong>" 50)
where
"f \<cong> g \<longleftrightarrow>
- (\<exists>c\<^isub>1>0. \<exists>c\<^isub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m)"
+ (\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m)"
text {*
- This yields @{text "f \<cong> g \<longleftrightarrow> f \<in> \<Theta>(g)"}. Concerning @{text "c\<^isub>1"} and @{text "c\<^isub>2"}
+ This yields @{text "f \<cong> g \<longleftrightarrow> f \<in> \<Theta>(g)"}. Concerning @{text "c\<^sub>1"} and @{text "c\<^sub>2"}
restricted to @{typ nat}, see note above on @{text "(\<lesssim>)"}.
*}
lemma equiv_funI [intro?]:
- assumes "\<exists>c\<^isub>1>0. \<exists>c\<^isub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m"
+ assumes "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
shows "f \<cong> g"
unfolding equiv_fun_def by (rule assms)
lemma not_equiv_funI:
- assumes "\<And>c\<^isub>1 c\<^isub>2 n. c\<^isub>1 > 0 \<Longrightarrow> c\<^isub>2 > 0 \<Longrightarrow>
- \<exists>m>n. c\<^isub>1 * f m < g m \<or> c\<^isub>2 * g m < f m"
+ assumes "\<And>c\<^sub>1 c\<^sub>2 n. c\<^sub>1 > 0 \<Longrightarrow> c\<^sub>2 > 0 \<Longrightarrow>
+ \<exists>m>n. c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
shows "\<not> f \<cong> g"
using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
lemma equiv_funE [elim?]:
assumes "f \<cong> g"
- obtains n c\<^isub>1 c\<^isub>2 where "c\<^isub>1 > 0" and "c\<^isub>2 > 0"
- and "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m"
+ obtains n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
+ and "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
using assms unfolding equiv_fun_def by blast
lemma not_equiv_funE:
- fixes n c\<^isub>1 c\<^isub>2
- assumes "\<not> f \<cong> g" and "c\<^isub>1 > 0" and "c\<^isub>2 > 0"
+ fixes n c\<^sub>1 c\<^sub>2
+ assumes "\<not> f \<cong> g" and "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
obtains m where "m > n"
- and "c\<^isub>1 * f m < g m \<or> c\<^isub>2 * g m < f m"
+ and "c\<^sub>1 * f m < g m \<or> c\<^sub>2 * g m < f m"
using assms unfolding equiv_fun_def linorder_not_le [symmetric] by blast
@@ -207,23 +207,23 @@
assume "f \<lesssim> g" and "g \<lesssim> h"
show "f \<lesssim> h"
proof
- from `f \<lesssim> g` obtain n\<^isub>1 c\<^isub>1
- where "c\<^isub>1 > 0" and P\<^isub>1: "\<And>m. m > n\<^isub>1 \<Longrightarrow> f m \<le> c\<^isub>1 * g m"
+ from `f \<lesssim> g` obtain n\<^sub>1 c\<^sub>1
+ where "c\<^sub>1 > 0" and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m"
by rule blast
- from `g \<lesssim> h` obtain n\<^isub>2 c\<^isub>2
- where "c\<^isub>2 > 0" and P\<^isub>2: "\<And>m. m > n\<^isub>2 \<Longrightarrow> g m \<le> c\<^isub>2 * h m"
+ from `g \<lesssim> h` obtain n\<^sub>2 c\<^sub>2
+ where "c\<^sub>2 > 0" and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * h m"
by rule blast
- have "\<forall>m>max n\<^isub>1 n\<^isub>2. f m \<le> (c\<^isub>1 * c\<^isub>2) * h m"
+ have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m"
proof (rule allI, rule impI)
fix m
- assume Q: "m > max n\<^isub>1 n\<^isub>2"
- from P\<^isub>1 Q have *: "f m \<le> c\<^isub>1 * g m" by simp
- from P\<^isub>2 Q have "g m \<le> c\<^isub>2 * h m" by simp
- with `c\<^isub>1 > 0` have "c\<^isub>1 * g m \<le> (c\<^isub>1 * c\<^isub>2) * h m" by simp
- with * show "f m \<le> (c\<^isub>1 * c\<^isub>2) * h m" by (rule order_trans)
+ assume Q: "m > max n\<^sub>1 n\<^sub>2"
+ from P\<^sub>1 Q have *: "f m \<le> c\<^sub>1 * g m" by simp
+ from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * h m" by simp
+ with `c\<^sub>1 > 0` have "c\<^sub>1 * g m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by simp
+ with * show "f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by (rule order_trans)
qed
- then have "\<exists>n. \<forall>m>n. f m \<le> (c\<^isub>1 * c\<^isub>2) * h m" by rule
- moreover from `c\<^isub>1 > 0` `c\<^isub>2 > 0` have "c\<^isub>1 * c\<^isub>2 > 0" by (rule mult_pos_pos)
+ then have "\<exists>n. \<forall>m>n. f m \<le> (c\<^sub>1 * c\<^sub>2) * h m" by rule
+ moreover from `c\<^sub>1 > 0` `c\<^sub>2 > 0` have "c\<^sub>1 * c\<^sub>2 > 0" by (rule mult_pos_pos)
ultimately show "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * h m" by blast
qed
qed
@@ -234,43 +234,43 @@
show "f \<lesssim> g \<and> g \<lesssim> f \<longleftrightarrow> f \<cong> g"
proof
assume "f \<cong> g"
- then obtain n c\<^isub>1 c\<^isub>2 where "c\<^isub>1 > 0" and "c\<^isub>2 > 0"
- and *: "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m"
+ then obtain n c\<^sub>1 c\<^sub>2 where "c\<^sub>1 > 0" and "c\<^sub>2 > 0"
+ and *: "\<And>m. m > n \<Longrightarrow> f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
by rule blast
- have "\<forall>m>n. f m \<le> c\<^isub>1 * g m"
+ have "\<forall>m>n. f m \<le> c\<^sub>1 * g m"
proof (rule allI, rule impI)
fix m
assume "m > n"
- with * show "f m \<le> c\<^isub>1 * g m" by simp
+ with * show "f m \<le> c\<^sub>1 * g m" by simp
qed
- with `c\<^isub>1 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
+ with `c\<^sub>1 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m" by blast
then have "f \<lesssim> g" ..
- have "\<forall>m>n. g m \<le> c\<^isub>2 * f m"
+ have "\<forall>m>n. g m \<le> c\<^sub>2 * f m"
proof (rule allI, rule impI)
fix m
assume "m > n"
- with * show "g m \<le> c\<^isub>2 * f m" by simp
+ with * show "g m \<le> c\<^sub>2 * f m" by simp
qed
- with `c\<^isub>2 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
+ with `c\<^sub>2 > 0` have "\<exists>c>0. \<exists>n. \<forall>m>n. g m \<le> c * f m" by blast
then have "g \<lesssim> f" ..
from `f \<lesssim> g` and `g \<lesssim> f` show "f \<lesssim> g \<and> g \<lesssim> f" ..
next
assume "f \<lesssim> g \<and> g \<lesssim> f"
then have "f \<lesssim> g" and "g \<lesssim> f" by auto
- from `f \<lesssim> g` obtain n\<^isub>1 c\<^isub>1 where "c\<^isub>1 > 0"
- and P\<^isub>1: "\<And>m. m > n\<^isub>1 \<Longrightarrow> f m \<le> c\<^isub>1 * g m" by rule blast
- from `g \<lesssim> f` obtain n\<^isub>2 c\<^isub>2 where "c\<^isub>2 > 0"
- and P\<^isub>2: "\<And>m. m > n\<^isub>2 \<Longrightarrow> g m \<le> c\<^isub>2 * f m" by rule blast
- have "\<forall>m>max n\<^isub>1 n\<^isub>2. f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m"
+ from `f \<lesssim> g` obtain n\<^sub>1 c\<^sub>1 where "c\<^sub>1 > 0"
+ and P\<^sub>1: "\<And>m. m > n\<^sub>1 \<Longrightarrow> f m \<le> c\<^sub>1 * g m" by rule blast
+ from `g \<lesssim> f` obtain n\<^sub>2 c\<^sub>2 where "c\<^sub>2 > 0"
+ and P\<^sub>2: "\<And>m. m > n\<^sub>2 \<Longrightarrow> g m \<le> c\<^sub>2 * f m" by rule blast
+ have "\<forall>m>max n\<^sub>1 n\<^sub>2. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m"
proof (rule allI, rule impI)
fix m
- assume Q: "m > max n\<^isub>1 n\<^isub>2"
- from P\<^isub>1 Q have "f m \<le> c\<^isub>1 * g m" by simp
- moreover from P\<^isub>2 Q have "g m \<le> c\<^isub>2 * f m" by simp
- ultimately show "f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m" ..
+ assume Q: "m > max n\<^sub>1 n\<^sub>2"
+ from P\<^sub>1 Q have "f m \<le> c\<^sub>1 * g m" by simp
+ moreover from P\<^sub>2 Q have "g m \<le> c\<^sub>2 * f m" by simp
+ ultimately show "f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" ..
qed
- with `c\<^isub>1 > 0` `c\<^isub>2 > 0` have "\<exists>c\<^isub>1>0. \<exists>c\<^isub>2>0. \<exists>n.
- \<forall>m>n. f m \<le> c\<^isub>1 * g m \<and> g m \<le> c\<^isub>2 * f m" by blast
+ with `c\<^sub>1 > 0` `c\<^sub>2 > 0` have "\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n.
+ \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m" by blast
then show "f \<cong> g" ..
qed
qed
@@ -318,15 +318,15 @@
proof (rule less_fun_strongI)
fix c :: nat
assume "0 < c"
- have "\<forall>m>(Suc c)\<twosuperior>. c * Discrete.sqrt m < id m"
+ have "\<forall>m>(Suc c)\<^sup>2. c * Discrete.sqrt m < id m"
proof (rule allI, rule impI)
fix m
- assume "(Suc c)\<twosuperior> < m"
- then have "(Suc c)\<twosuperior> \<le> m" by simp
- with mono_sqrt have "Discrete.sqrt ((Suc c)\<twosuperior>) \<le> Discrete.sqrt m" by (rule monoE)
+ assume "(Suc c)\<^sup>2 < m"
+ then have "(Suc c)\<^sup>2 \<le> m" by simp
+ with mono_sqrt have "Discrete.sqrt ((Suc c)\<^sup>2) \<le> Discrete.sqrt m" by (rule monoE)
then have "Suc c \<le> Discrete.sqrt m" by simp
then have "c < Discrete.sqrt m" by simp
- moreover from `(Suc c)\<twosuperior> < m` have "Discrete.sqrt m > 0" by simp
+ moreover from `(Suc c)\<^sup>2 < m` have "Discrete.sqrt m > 0" by simp
ultimately have "c * Discrete.sqrt m < Discrete.sqrt m * Discrete.sqrt m" by simp
also have "\<dots> \<le> m" by (simp add: power2_eq_square [symmetric])
finally show "c * Discrete.sqrt m < id m" by simp
@@ -334,7 +334,7 @@
then show "\<exists>n. \<forall>m>n. c * Discrete.sqrt m < id m" ..
qed
-lemma "id \<prec> (\<lambda>n. n\<twosuperior>)"
+lemma "id \<prec> (\<lambda>n. n\<^sup>2)"
by (rule less_fun_strongI) (auto simp add: power2_eq_square)
lemma "(\<lambda>n. n ^ k) \<prec> (\<lambda>n. n ^ Suc k)"
--- a/src/HOL/Library/Inner_Product.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Inner_Product.thy Tue Aug 13 16:25:47 2013 +0200
@@ -78,11 +78,11 @@
lemma inner_gt_zero_iff [simp]: "0 < inner x x \<longleftrightarrow> x \<noteq> 0"
by (simp add: order_less_le)
-lemma power2_norm_eq_inner: "(norm x)\<twosuperior> = inner x x"
+lemma power2_norm_eq_inner: "(norm x)\<^sup>2 = inner x x"
by (simp add: norm_eq_sqrt_inner)
lemma Cauchy_Schwarz_ineq:
- "(inner x y)\<twosuperior> \<le> inner x x * inner y y"
+ "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
proof (cases)
assume "y = 0"
thus ?thesis by simp
@@ -93,21 +93,21 @@
by (rule inner_ge_zero)
also have "\<dots> = inner x x - inner y x * ?r"
by (simp add: inner_diff)
- also have "\<dots> = inner x x - (inner x y)\<twosuperior> / inner y y"
+ also have "\<dots> = inner x x - (inner x y)\<^sup>2 / inner y y"
by (simp add: power2_eq_square inner_commute)
- finally have "0 \<le> inner x x - (inner x y)\<twosuperior> / inner y y" .
- hence "(inner x y)\<twosuperior> / inner y y \<le> inner x x"
+ finally have "0 \<le> inner x x - (inner x y)\<^sup>2 / inner y y" .
+ hence "(inner x y)\<^sup>2 / inner y y \<le> inner x x"
by (simp add: le_diff_eq)
- thus "(inner x y)\<twosuperior> \<le> inner x x * inner y y"
+ thus "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
by (simp add: pos_divide_le_eq y)
qed
lemma Cauchy_Schwarz_ineq2:
"\<bar>inner x y\<bar> \<le> norm x * norm y"
proof (rule power2_le_imp_le)
- have "(inner x y)\<twosuperior> \<le> inner x x * inner y y"
+ have "(inner x y)\<^sup>2 \<le> inner x x * inner y y"
using Cauchy_Schwarz_ineq .
- thus "\<bar>inner x y\<bar>\<twosuperior> \<le> (norm x * norm y)\<twosuperior>"
+ thus "\<bar>inner x y\<bar>\<^sup>2 \<le> (norm x * norm y)\<^sup>2"
by (simp add: power_mult_distrib power2_norm_eq_inner)
show "0 \<le> norm x * norm y"
unfolding norm_eq_sqrt_inner
@@ -123,13 +123,13 @@
proof (rule power2_le_imp_le)
have "inner x y \<le> norm x * norm y"
by (rule order_trans [OF abs_ge_self Cauchy_Schwarz_ineq2])
- thus "(norm (x + y))\<twosuperior> \<le> (norm x + norm y)\<twosuperior>"
+ thus "(norm (x + y))\<^sup>2 \<le> (norm x + norm y)\<^sup>2"
unfolding power2_sum power2_norm_eq_inner
by (simp add: inner_add inner_commute)
show "0 \<le> norm x + norm y"
unfolding norm_eq_sqrt_inner by simp
qed
- have "sqrt (a\<twosuperior> * inner x x) = \<bar>a\<bar> * sqrt (inner x x)"
+ have "sqrt (a\<^sup>2 * inner x x) = \<bar>a\<bar> * sqrt (inner x x)"
by (simp add: real_sqrt_mult_distrib)
then show "norm (a *\<^sub>R x) = \<bar>a\<bar> * norm x"
unfolding norm_eq_sqrt_inner
@@ -324,7 +324,7 @@
lemma GDERIV_inverse:
"\<lbrakk>GDERIV f x :> df; f x \<noteq> 0\<rbrakk>
- \<Longrightarrow> GDERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x))\<twosuperior> *\<^sub>R df"
+ \<Longrightarrow> GDERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x))\<^sup>2 *\<^sub>R df"
apply (erule GDERIV_DERIV_compose)
apply (erule DERIV_inverse [folded numeral_2_eq_2])
done
--- a/src/HOL/Library/Kleene_Algebra.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Kleene_Algebra.thy Tue Aug 13 16:25:47 2013 +0200
@@ -321,10 +321,10 @@
lemma ka25: "star y * star x \<le> star x * star y \<Longrightarrow> star (star y * star x) \<le> star x * star y"
proof -
assume "star y * star x \<le> star x * star y"
- hence "\<forall>x\<^isub>1. star y * (star x * x\<^isub>1) \<le> star x * (star y * x\<^isub>1)" by (metis mult_assoc mult_right_mono zero_minimum)
+ hence "\<forall>x\<^sub>1. star y * (star x * x\<^sub>1) \<le> star x * (star y * x\<^sub>1)" by (metis mult_assoc mult_right_mono zero_minimum)
hence "star y * (star x * star y) \<le> star x * star y" by (metis star_mult_idem)
- hence "\<exists>x\<^isub>1. star (star y * star x) * star x\<^isub>1 \<le> star x * star y" by (metis star_decomp star_idemp star_simulation_leq_2 star_slide)
- hence "\<exists>x\<^isub>1\<ge>star (star y * star x). x\<^isub>1 \<le> star x * star y" by (metis x_less_star)
+ hence "\<exists>x\<^sub>1. star (star y * star x) * star x\<^sub>1 \<le> star x * star y" by (metis star_decomp star_idemp star_simulation_leq_2 star_slide)
+ hence "\<exists>x\<^sub>1\<ge>star (star y * star x). x\<^sub>1 \<le> star x * star y" by (metis x_less_star)
thus "star (star y * star x) \<le> star x * star y" by (metis order_trans)
qed
--- a/src/HOL/Library/Product_Lexorder.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Product_Lexorder.thy Tue Aug 13 16:25:47 2013 +0200
@@ -94,23 +94,23 @@
case (Pair a b)
show "P (a, b)"
proof (induct a arbitrary: b rule: less_induct)
- case (less a\<^isub>1) note a\<^isub>1 = this
- show "P (a\<^isub>1, b)"
+ case (less a\<^sub>1) note a\<^sub>1 = this
+ show "P (a\<^sub>1, b)"
proof (induct b rule: less_induct)
- case (less b\<^isub>1) note b\<^isub>1 = this
- show "P (a\<^isub>1, b\<^isub>1)"
+ case (less b\<^sub>1) note b\<^sub>1 = this
+ show "P (a\<^sub>1, b\<^sub>1)"
proof (rule P)
- fix p assume p: "p < (a\<^isub>1, b\<^isub>1)"
+ fix p assume p: "p < (a\<^sub>1, b\<^sub>1)"
show "P p"
- proof (cases "fst p < a\<^isub>1")
+ proof (cases "fst p < a\<^sub>1")
case True
- then have "P (fst p, snd p)" by (rule a\<^isub>1)
+ then have "P (fst p, snd p)" by (rule a\<^sub>1)
then show ?thesis by simp
next
case False
- with p have 1: "a\<^isub>1 = fst p" and 2: "snd p < b\<^isub>1"
+ with p have 1: "a\<^sub>1 = fst p" and 2: "snd p < b\<^sub>1"
by (simp_all add: less_prod_def')
- from 2 have "P (a\<^isub>1, snd p)" by (rule b\<^isub>1)
+ from 2 have "P (a\<^sub>1, snd p)" by (rule b\<^sub>1)
with 1 show ?thesis by simp
qed
qed
--- a/src/HOL/Library/Product_Vector.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Product_Vector.thy Tue Aug 13 16:25:47 2013 +0200
@@ -275,9 +275,9 @@
begin
definition dist_prod_def:
- "dist x y = sqrt ((dist (fst x) (fst y))\<twosuperior> + (dist (snd x) (snd y))\<twosuperior>)"
+ "dist x y = sqrt ((dist (fst x) (fst y))\<^sup>2 + (dist (snd x) (snd y))\<^sup>2)"
-lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<twosuperior> + (dist b d)\<twosuperior>)"
+lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<^sup>2 + (dist b d)\<^sup>2)"
unfolding dist_prod_def by simp
lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y"
@@ -335,7 +335,7 @@
def r \<equiv> "e / sqrt 2" and s \<equiv> "e / sqrt 2"
from `0 < e` have "0 < r" and "0 < s"
unfolding r_def s_def by (simp_all add: divide_pos_pos)
- from `0 < e` have "e = sqrt (r\<twosuperior> + s\<twosuperior>)"
+ from `0 < e` have "e = sqrt (r\<^sup>2 + s\<^sup>2)"
unfolding r_def s_def by (simp add: power_divide)
def A \<equiv> "{y. dist (fst x) y < r}" and B \<equiv> "{y. dist (snd x) y < s}"
have "open A" and "open B"
@@ -349,7 +349,7 @@
hence "dist a (fst x) < r" and "dist b (snd x) < s"
unfolding A_def B_def by (simp_all add: dist_commute)
hence "dist (a, b) x < e"
- unfolding dist_prod_def `e = sqrt (r\<twosuperior> + s\<twosuperior>)`
+ unfolding dist_prod_def `e = sqrt (r\<^sup>2 + s\<^sup>2)`
by (simp add: add_strict_mono power_strict_mono)
thus "(a, b) \<in> S"
by (simp add: S)
@@ -406,12 +406,12 @@
begin
definition norm_prod_def:
- "norm x = sqrt ((norm (fst x))\<twosuperior> + (norm (snd x))\<twosuperior>)"
+ "norm x = sqrt ((norm (fst x))\<^sup>2 + (norm (snd x))\<^sup>2)"
definition sgn_prod_def:
"sgn (x::'a \<times> 'b) = scaleR (inverse (norm x)) x"
-lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\<twosuperior> + (norm b)\<twosuperior>)"
+lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\<^sup>2 + (norm b)\<^sup>2)"
unfolding norm_prod_def by simp
instance proof
--- a/src/HOL/Library/Sublist.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Library/Sublist.thy Tue Aug 13 16:25:47 2013 +0200
@@ -115,7 +115,7 @@
by (auto simp add: prefixeq_def)
lemma prefixeq_same_cases:
- "prefixeq (xs\<^isub>1::'a list) ys \<Longrightarrow> prefixeq xs\<^isub>2 ys \<Longrightarrow> prefixeq xs\<^isub>1 xs\<^isub>2 \<or> prefixeq xs\<^isub>2 xs\<^isub>1"
+ "prefixeq (xs\<^sub>1::'a list) ys \<Longrightarrow> prefixeq xs\<^sub>2 ys \<Longrightarrow> prefixeq xs\<^sub>1 xs\<^sub>2 \<or> prefixeq xs\<^sub>2 xs\<^sub>1"
unfolding prefixeq_def by (metis append_eq_append_conv2)
lemma set_mono_prefixeq: "prefixeq xs ys \<Longrightarrow> set xs \<subseteq> set ys"
--- a/src/HOL/List.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/List.thy Tue Aug 13 16:25:47 2013 +0200
@@ -2095,13 +2095,13 @@
done
lemma append_eq_append_conv_if:
- "(xs\<^isub>1 @ xs\<^isub>2 = ys\<^isub>1 @ ys\<^isub>2) =
- (if size xs\<^isub>1 \<le> size ys\<^isub>1
- then xs\<^isub>1 = take (size xs\<^isub>1) ys\<^isub>1 \<and> xs\<^isub>2 = drop (size xs\<^isub>1) ys\<^isub>1 @ ys\<^isub>2
- else take (size ys\<^isub>1) xs\<^isub>1 = ys\<^isub>1 \<and> drop (size ys\<^isub>1) xs\<^isub>1 @ xs\<^isub>2 = ys\<^isub>2)"
-apply(induct xs\<^isub>1 arbitrary: ys\<^isub>1)
+ "(xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>1 @ ys\<^sub>2) =
+ (if size xs\<^sub>1 \<le> size ys\<^sub>1
+ then xs\<^sub>1 = take (size xs\<^sub>1) ys\<^sub>1 \<and> xs\<^sub>2 = drop (size xs\<^sub>1) ys\<^sub>1 @ ys\<^sub>2
+ else take (size ys\<^sub>1) xs\<^sub>1 = ys\<^sub>1 \<and> drop (size ys\<^sub>1) xs\<^sub>1 @ xs\<^sub>2 = ys\<^sub>2)"
+apply(induct xs\<^sub>1 arbitrary: ys\<^sub>1)
apply simp
-apply(case_tac ys\<^isub>1)
+apply(case_tac ys\<^sub>1)
apply simp_all
done
--- a/src/HOL/Map.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Map.thy Tue Aug 13 16:25:47 2013 +0200
@@ -48,7 +48,7 @@
definition
map_le :: "('a ~=> 'b) => ('a ~=> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50) where
- "(m\<^isub>1 \<subseteq>\<^sub>m m\<^isub>2) = (\<forall>a \<in> dom m\<^isub>1. m\<^isub>1 a = m\<^isub>2 a)"
+ "(m\<^sub>1 \<subseteq>\<^sub>m m\<^sub>2) = (\<forall>a \<in> dom m\<^sub>1. m\<^sub>1 a = m\<^sub>2 a)"
nonterminal maplets and maplet
--- a/src/HOL/Metis_Examples/Big_O.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy Tue Aug 13 16:25:47 2013 +0200
@@ -42,20 +42,20 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> \<bar>x\<^isub>1\<bar>" by (metis abs_ge_zero)
- have F2: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1)
- have F3: "\<forall>x\<^isub>1 x\<^isub>3. x\<^isub>3 \<le> \<bar>h x\<^isub>1\<bar> \<longrightarrow> x\<^isub>3 \<le> c * \<bar>f x\<^isub>1\<bar>" by (metis A1 order_trans)
- have F4: "\<forall>x\<^isub>2 x\<^isub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^isub>3\<bar> * \<bar>x\<^isub>2\<bar> = \<bar>x\<^isub>3 * x\<^isub>2\<bar>"
+ have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> \<bar>x\<^sub>1\<bar>" by (metis abs_ge_zero)
+ have F2: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F3: "\<forall>x\<^sub>1 x\<^sub>3. x\<^sub>3 \<le> \<bar>h x\<^sub>1\<bar> \<longrightarrow> x\<^sub>3 \<le> c * \<bar>f x\<^sub>1\<bar>" by (metis A1 order_trans)
+ have F4: "\<forall>x\<^sub>2 x\<^sub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
by (metis abs_mult)
- have F5: "\<forall>x\<^isub>3 x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^isub>1 \<longrightarrow> \<bar>x\<^isub>3 * x\<^isub>1\<bar> = \<bar>x\<^isub>3\<bar> * x\<^isub>1"
+ have F5: "\<forall>x\<^sub>3 x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1"
by (metis abs_mult_pos)
- hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = \<bar>1\<bar> * x\<^isub>1" by (metis F2)
- hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F2 abs_one)
- hence "\<forall>x\<^isub>3. 0 \<le> \<bar>h x\<^isub>3\<bar> \<longrightarrow> \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis F3)
- hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis F1)
- hence "\<forall>x\<^isub>3. (0\<Colon>'a) \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^isub>3\<bar>" by (metis F5)
- hence "\<forall>x\<^isub>3. (0\<Colon>'a) \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F4)
- hence "\<forall>x\<^isub>3. c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F1)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = \<bar>1\<bar> * x\<^sub>1" by (metis F2)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F2 abs_one)
+ hence "\<forall>x\<^sub>3. 0 \<le> \<bar>h x\<^sub>3\<bar> \<longrightarrow> \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F3)
+ hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis F1)
+ hence "\<forall>x\<^sub>3. (0\<Colon>'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F5)
+ hence "\<forall>x\<^sub>3. (0\<Colon>'a) \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F4)
+ hence "\<forall>x\<^sub>3. c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F1)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F4)
qed
@@ -73,12 +73,12 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1)
- have F2: "\<forall>x\<^isub>2 x\<^isub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^isub>3\<bar> * \<bar>x\<^isub>2\<bar> = \<bar>x\<^isub>3 * x\<^isub>2\<bar>"
+ have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F2: "\<forall>x\<^sub>2 x\<^sub>3\<Colon>'a\<Colon>linordered_idom. \<bar>x\<^sub>3\<bar> * \<bar>x\<^sub>2\<bar> = \<bar>x\<^sub>3 * x\<^sub>2\<bar>"
by (metis abs_mult)
- have "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F1 abs_mult_pos abs_one)
- hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>" by (metis A1 abs_ge_zero order_trans)
- hence "\<forall>x\<^isub>3. 0 \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c * f x\<^isub>3\<bar>" by (metis F2 abs_mult_pos)
+ have "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_mult_pos abs_one)
+ hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>" by (metis A1 abs_ge_zero order_trans)
+ hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c * f x\<^sub>3\<bar>" by (metis F2 abs_mult_pos)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis F2)
qed
@@ -96,10 +96,10 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have F1: "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1)
- have F2: "\<forall>x\<^isub>3 x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^isub>1 \<longrightarrow> \<bar>x\<^isub>3 * x\<^isub>1\<bar> = \<bar>x\<^isub>3\<bar> * x\<^isub>1" by (metis abs_mult_pos)
- hence "\<forall>x\<^isub>1\<ge>0. \<bar>x\<^isub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^isub>1" by (metis F1 abs_one)
- hence "\<forall>x\<^isub>3. 0 \<le> \<bar>f x\<^isub>3\<bar> \<longrightarrow> c * \<bar>f x\<^isub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^isub>3\<bar>" by (metis F2 A1 abs_ge_zero order_trans)
+ have F1: "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ have F2: "\<forall>x\<^sub>3 x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 0 \<le> x\<^sub>1 \<longrightarrow> \<bar>x\<^sub>3 * x\<^sub>1\<bar> = \<bar>x\<^sub>3\<bar> * x\<^sub>1" by (metis abs_mult_pos)
+ hence "\<forall>x\<^sub>1\<ge>0. \<bar>x\<^sub>1\<Colon>'a\<Colon>linordered_idom\<bar> = x\<^sub>1" by (metis F1 abs_one)
+ hence "\<forall>x\<^sub>3. 0 \<le> \<bar>f x\<^sub>3\<bar> \<longrightarrow> c * \<bar>f x\<^sub>3\<bar> = \<bar>c\<bar> * \<bar>f x\<^sub>3\<bar>" by (metis F2 A1 abs_ge_zero order_trans)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis A1 abs_ge_zero)
qed
@@ -116,8 +116,8 @@
proof -
fix c :: 'a and x :: 'b
assume A1: "\<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>"
- have "\<forall>x\<^isub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^isub>1 = x\<^isub>1" by (metis mult_1)
- hence "\<forall>x\<^isub>3. \<bar>c * \<bar>f x\<^isub>3\<bar>\<bar> = c * \<bar>f x\<^isub>3\<bar>"
+ have "\<forall>x\<^sub>1\<Colon>'a\<Colon>linordered_idom. 1 * x\<^sub>1 = x\<^sub>1" by (metis mult_1)
+ hence "\<forall>x\<^sub>3. \<bar>c * \<bar>f x\<^sub>3\<bar>\<bar> = c * \<bar>f x\<^sub>3\<bar>"
by (metis A1 abs_ge_zero order_trans abs_mult_pos abs_one)
hence "\<bar>h x\<bar> \<le> \<bar>c * f x\<bar>" by (metis A1 abs_ge_zero abs_mult_pos abs_mult)
thus "\<bar>h x\<bar> \<le> \<bar>c\<bar> * \<bar>f x\<bar>" by (metis abs_mult)
@@ -459,11 +459,11 @@
have "(0\<Colon>'a) < \<bar>c\<bar> \<longrightarrow> (0\<Colon>'a) < \<bar>inverse c\<bar>" using F1 by (metis positive_imp_inverse_positive)
hence "(0\<Colon>'a) < \<bar>inverse c\<bar>" using F2 by metis
hence F3: "(0\<Colon>'a) \<le> \<bar>inverse c\<bar>" by (metis order_le_less)
- have "\<exists>(u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar>"
+ have "\<exists>(u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar>"
using A2 by metis
- hence F4: "\<exists>(u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^isub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0\<Colon>'a) \<le> \<bar>inverse c\<bar>"
+ hence F4: "\<exists>(u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>g (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<le> u * \<bar>f (SKF\<^sub>7 (\<bar>inverse c\<bar> * u))\<bar> \<and> (0\<Colon>'a) \<le> \<bar>inverse c\<bar>"
using F3 by metis
- hence "\<exists>(v\<Colon>'a) (u\<Colon>'a) SKF\<^isub>7\<Colon>'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^isub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^isub>7 (u * v))\<bar>)"
+ hence "\<exists>(v\<Colon>'a) (u\<Colon>'a) SKF\<^sub>7\<Colon>'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^sub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^sub>7 (u * v))\<bar>)"
by (metis comm_mult_left_mono)
thus "\<exists>ca\<Colon>'a. \<forall>x\<Colon>'b. \<bar>inverse c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>"
using A2 F4 by (metis ab_semigroup_mult_class.mult_ac(1) comm_mult_left_mono)
--- a/src/HOL/Metis_Examples/Binary_Tree.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Metis_Examples/Binary_Tree.thy Tue Aug 13 16:25:47 2013 +0200
@@ -59,9 +59,9 @@
proof (induct t)
case Lf thus ?case
proof -
- let "?p\<^isub>1 x\<^isub>1" = "x\<^isub>1 \<noteq> n_leaves (reflect (Lf::'a bt))"
- have "\<not> ?p\<^isub>1 (Suc 0)" by (metis reflect.simps(1) n_leaves.simps(1))
- hence "\<not> ?p\<^isub>1 (n_leaves (Lf::'a bt))" by (metis n_leaves.simps(1))
+ let "?p\<^sub>1 x\<^sub>1" = "x\<^sub>1 \<noteq> n_leaves (reflect (Lf::'a bt))"
+ have "\<not> ?p\<^sub>1 (Suc 0)" by (metis reflect.simps(1) n_leaves.simps(1))
+ hence "\<not> ?p\<^sub>1 (n_leaves (Lf::'a bt))" by (metis n_leaves.simps(1))
thus "n_leaves (reflect (Lf::'a bt)) = n_leaves (Lf::'a bt)" by metis
qed
next
--- a/src/HOL/Metis_Examples/Message.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Metis_Examples/Message.thy Tue Aug 13 16:25:47 2013 +0200
@@ -677,10 +677,10 @@
lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
proof -
- have "\<forall>x\<^isub>2 x\<^isub>1. synth x\<^isub>1 \<union> analz (x\<^isub>1 \<union> x\<^isub>2) = analz (synth x\<^isub>1 \<union> x\<^isub>2)" by (metis Un_commute analz_synth_Un)
- hence "\<forall>x\<^isub>1. synth x\<^isub>1 \<union> analz x\<^isub>1 = analz (synth x\<^isub>1 \<union> {})" by (metis Un_empty_right)
- hence "\<forall>x\<^isub>1. synth x\<^isub>1 \<union> analz x\<^isub>1 = analz (synth x\<^isub>1)" by (metis Un_empty_right)
- hence "\<forall>x\<^isub>1. analz x\<^isub>1 \<union> synth x\<^isub>1 = analz (synth x\<^isub>1)" by (metis Un_commute)
+ have "\<forall>x\<^sub>2 x\<^sub>1. synth x\<^sub>1 \<union> analz (x\<^sub>1 \<union> x\<^sub>2) = analz (synth x\<^sub>1 \<union> x\<^sub>2)" by (metis Un_commute analz_synth_Un)
+ hence "\<forall>x\<^sub>1. synth x\<^sub>1 \<union> analz x\<^sub>1 = analz (synth x\<^sub>1 \<union> {})" by (metis Un_empty_right)
+ hence "\<forall>x\<^sub>1. synth x\<^sub>1 \<union> analz x\<^sub>1 = analz (synth x\<^sub>1)" by (metis Un_empty_right)
+ hence "\<forall>x\<^sub>1. analz x\<^sub>1 \<union> synth x\<^sub>1 = analz (synth x\<^sub>1)" by (metis Un_commute)
thus "analz (synth H) = analz H \<union> synth H" by metis
qed
@@ -690,8 +690,8 @@
lemma parts_insert_subset_Un: "X \<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
proof -
assume "X \<in> G"
- hence "\<forall>x\<^isub>1. G \<subseteq> x\<^isub>1 \<longrightarrow> X \<in> x\<^isub>1 " by auto
- hence "\<forall>x\<^isub>1. X \<in> G \<union> x\<^isub>1" by (metis Un_upper1)
+ hence "\<forall>x\<^sub>1. G \<subseteq> x\<^sub>1 \<longrightarrow> X \<in> x\<^sub>1 " by auto
+ hence "\<forall>x\<^sub>1. X \<in> G \<union> x\<^sub>1" by (metis Un_upper1)
hence "insert X H \<subseteq> G \<union> H" by (metis Un_upper2 insert_subset)
hence "parts (insert X H) \<subseteq> parts (G \<union> H)" by (metis parts_mono)
thus "parts (insert X H) \<subseteq> parts G \<union> parts H" by (metis parts_Un)
@@ -702,22 +702,22 @@
parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
proof -
assume A1: "X \<in> synth (analz H)"
- have F1: "\<forall>x\<^isub>1. analz x\<^isub>1 \<union> synth (analz x\<^isub>1) = analz (synth (analz x\<^isub>1))"
+ have F1: "\<forall>x\<^sub>1. analz x\<^sub>1 \<union> synth (analz x\<^sub>1) = analz (synth (analz x\<^sub>1))"
by (metis analz_idem analz_synth)
- have F2: "\<forall>x\<^isub>1. parts x\<^isub>1 \<union> synth (analz x\<^isub>1) = parts (synth (analz x\<^isub>1))"
+ have F2: "\<forall>x\<^sub>1. parts x\<^sub>1 \<union> synth (analz x\<^sub>1) = parts (synth (analz x\<^sub>1))"
by (metis parts_analz parts_synth)
have F3: "X \<in> synth (analz H)" using A1 by metis
- have "\<forall>x\<^isub>2 x\<^isub>1\<Colon>msg set. x\<^isub>1 \<le> sup x\<^isub>1 x\<^isub>2" by (metis inf_sup_ord(3))
- hence F4: "\<forall>x\<^isub>1. analz x\<^isub>1 \<subseteq> analz (synth x\<^isub>1)" by (metis analz_synth)
+ have "\<forall>x\<^sub>2 x\<^sub>1\<Colon>msg set. x\<^sub>1 \<le> sup x\<^sub>1 x\<^sub>2" by (metis inf_sup_ord(3))
+ hence F4: "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> analz (synth x\<^sub>1)" by (metis analz_synth)
have F5: "X \<in> synth (analz H)" using F3 by metis
- have "\<forall>x\<^isub>1. analz x\<^isub>1 \<subseteq> synth (analz x\<^isub>1)
- \<longrightarrow> analz (synth (analz x\<^isub>1)) = synth (analz x\<^isub>1)"
+ have "\<forall>x\<^sub>1. analz x\<^sub>1 \<subseteq> synth (analz x\<^sub>1)
+ \<longrightarrow> analz (synth (analz x\<^sub>1)) = synth (analz x\<^sub>1)"
using F1 by (metis subset_Un_eq)
- hence F6: "\<forall>x\<^isub>1. analz (synth (analz x\<^isub>1)) = synth (analz x\<^isub>1)"
+ hence F6: "\<forall>x\<^sub>1. analz (synth (analz x\<^sub>1)) = synth (analz x\<^sub>1)"
by (metis synth_increasing)
- have "\<forall>x\<^isub>1. x\<^isub>1 \<subseteq> analz (synth x\<^isub>1)" using F4 by (metis analz_subset_iff)
- hence "\<forall>x\<^isub>1. x\<^isub>1 \<subseteq> analz (synth (analz x\<^isub>1))" by (metis analz_subset_iff)
- hence "\<forall>x\<^isub>1. x\<^isub>1 \<subseteq> synth (analz x\<^isub>1)" using F6 by metis
+ have "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> analz (synth x\<^sub>1)" using F4 by (metis analz_subset_iff)
+ hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> analz (synth (analz x\<^sub>1))" by (metis analz_subset_iff)
+ hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> synth (analz x\<^sub>1)" using F6 by metis
hence "H \<subseteq> synth (analz H)" by metis
hence "H \<subseteq> synth (analz H) \<and> X \<in> synth (analz H)" using F5 by metis
hence "insert X H \<subseteq> synth (analz H)" by (metis insert_subset)
--- a/src/HOL/Metis_Examples/Sets.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Metis_Examples/Sets.thy Tue Aug 13 16:25:47 2013 +0200
@@ -27,9 +27,9 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>1 \<union> x\<^isub>2" by (metis Un_commute Un_upper2)
- have F2a: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<longrightarrow> x\<^isub>2 = x\<^isub>2 \<union> x\<^isub>1" by (metis Un_commute subset_Un_eq)
- have F2: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis F2a subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>1 \<union> x\<^sub>2" by (metis Un_commute Un_upper2)
+ have F2a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
+ have F2: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F2a subset_Un_eq)
{ assume "\<not> Z \<subseteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
@@ -44,12 +44,12 @@
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)
- hence "\<exists>x\<^isub>1\<Colon>'a set. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1)
+ hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1) }
moreover
- { assume "\<exists>x\<^isub>1\<Colon>'a set. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
+ { assume "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume "\<not> Y \<subseteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1) }
moreover
@@ -60,7 +60,7 @@
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z" by (metis Un_subset_iff)
hence "Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> Y \<union> Z" by (metis F2)
- hence "\<exists>x\<^isub>1\<Colon>'a set. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1)
+ hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1)
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by blast }
@@ -75,14 +75,14 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis Un_commute subset_Un_eq)
- { assume AA1: "\<exists>x\<^isub>1\<Colon>'a set. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
+ have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
+ { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume AAA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
{ assume "\<not> Z \<subseteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
- hence "\<exists>x\<^isub>1\<Colon>'a set. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
+ hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AAA1 Un_subset_iff) }
moreover
@@ -101,7 +101,7 @@
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
moreover
{ assume "Y \<union> Z \<subseteq> X \<and> X \<noteq> Y \<union> Z"
- hence "\<exists>x\<^isub>1\<Colon>'a set. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
+ hence "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z" by (metis F1 Un_commute Un_upper2)
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
ultimately have "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_subset_iff) }
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by metis
@@ -112,12 +112,12 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1a: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<longrightarrow> x\<^isub>2 = x\<^isub>2 \<union> x\<^isub>1" by (metis Un_commute subset_Un_eq)
- have F1: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis F1a subset_Un_eq)
+ have F1a: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<longrightarrow> x\<^sub>2 = x\<^sub>2 \<union> x\<^sub>1" by (metis Un_commute subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis F1a subset_Un_eq)
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
moreover
- { assume AA1: "\<exists>x\<^isub>1\<Colon>'a set. (Z \<subseteq> x\<^isub>1 \<and> Y \<subseteq> x\<^isub>1) \<and> \<not> X \<subseteq> x\<^isub>1"
+ { assume AA1: "\<exists>x\<^sub>1\<Colon>'a set. (Z \<subseteq> x\<^sub>1 \<and> Y \<subseteq> x\<^sub>1) \<and> \<not> X \<subseteq> x\<^sub>1"
{ assume "(Z \<subseteq> X \<and> Y \<subseteq> X) \<and> Y \<union> Z \<noteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis F1 Un_commute Un_subset_iff Un_upper2) }
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 Un_commute Un_subset_iff Un_upper2) }
@@ -129,12 +129,12 @@
lemma (*equal_union: *)
"(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))"
proof -
- have F1: "\<forall>(x\<^isub>2\<Colon>'b set) x\<^isub>1\<Colon>'b set. x\<^isub>1 \<subseteq> x\<^isub>2 \<and> x\<^isub>2 \<subseteq> x\<^isub>1 \<longrightarrow> x\<^isub>1 = x\<^isub>2" by (metis Un_commute subset_Un_eq)
+ have F1: "\<forall>(x\<^sub>2\<Colon>'b set) x\<^sub>1\<Colon>'b set. x\<^sub>1 \<subseteq> x\<^sub>2 \<and> x\<^sub>2 \<subseteq> x\<^sub>1 \<longrightarrow> x\<^sub>1 = x\<^sub>2" by (metis Un_commute subset_Un_eq)
{ assume "\<not> Y \<subseteq> X"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_commute Un_upper2) }
moreover
{ assume AA1: "Y \<subseteq> X \<and> Y \<union> Z \<noteq> X"
- { assume "\<exists>x\<^isub>1\<Colon>'a set. Y \<subseteq> x\<^isub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^isub>1 \<union> Z"
+ { assume "\<exists>x\<^sub>1\<Colon>'a set. Y \<subseteq> x\<^sub>1 \<union> Z \<and> Y \<union> Z \<noteq> X \<and> \<not> X \<subseteq> x\<^sub>1 \<union> Z"
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_upper2) }
hence "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis AA1 F1 Un_commute Un_subset_iff Un_upper2) }
ultimately show "(X = Y \<union> Z) = (Y \<subseteq> X \<and> Z \<subseteq> X \<and> (\<forall>V\<Colon>'a set. Y \<subseteq> V \<and> Z \<subseteq> V \<longrightarrow> X \<subseteq> V))" by (metis Un_subset_iff Un_upper2)
@@ -170,10 +170,10 @@
"\<forall>x \<in> S. \<Union>S \<subseteq> x \<Longrightarrow> \<exists>z. S \<subseteq> {z}"
proof -
assume "\<forall>x \<in> S. \<Union>S \<subseteq> x"
- hence "\<forall>x\<^isub>1. x\<^isub>1 \<subseteq> \<Union>S \<and> x\<^isub>1 \<in> S \<longrightarrow> x\<^isub>1 = \<Union>S" by (metis set_eq_subset)
- hence "\<forall>x\<^isub>1. x\<^isub>1 \<in> S \<longrightarrow> x\<^isub>1 = \<Union>S" by (metis Union_upper)
- hence "\<forall>x\<^isub>1\<Colon>('a set) set. \<Union>S \<in> x\<^isub>1 \<longrightarrow> S \<subseteq> x\<^isub>1" by (metis subsetI)
- hence "\<forall>x\<^isub>1\<Colon>('a set) set. S \<subseteq> insert (\<Union>S) x\<^isub>1" by (metis insert_iff)
+ hence "\<forall>x\<^sub>1. x\<^sub>1 \<subseteq> \<Union>S \<and> x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis set_eq_subset)
+ hence "\<forall>x\<^sub>1. x\<^sub>1 \<in> S \<longrightarrow> x\<^sub>1 = \<Union>S" by (metis Union_upper)
+ hence "\<forall>x\<^sub>1\<Colon>('a set) set. \<Union>S \<in> x\<^sub>1 \<longrightarrow> S \<subseteq> x\<^sub>1" by (metis subsetI)
+ hence "\<forall>x\<^sub>1\<Colon>('a set) set. S \<subseteq> insert (\<Union>S) x\<^sub>1" by (metis insert_iff)
thus "\<exists>z. S \<subseteq> {z}" by metis
qed
--- a/src/HOL/Metis_Examples/Trans_Closure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Metis_Examples/Trans_Closure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -51,7 +51,7 @@
assume A3: "(a, b) \<in> R\<^sup>*"
assume A4: "(b, c) \<in> R\<^sup>*"
have "b \<noteq> c" using A1 A2 by metis
- hence "\<exists>x\<^isub>1. (b, x\<^isub>1) \<in> R" using A4 by (metis converse_rtranclE)
+ hence "\<exists>x\<^sub>1. (b, x\<^sub>1) \<in> R" using A4 by (metis converse_rtranclE)
thus "\<exists>c. (b, c) \<in> R \<and> (a, c) \<in> R\<^sup>*" using A3 by (metis transitive_closure_trans(6))
qed
--- a/src/HOL/MicroJava/DFA/Product.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/MicroJava/DFA/Product.thy Tue Aug 13 16:25:47 2013 +0200
@@ -47,7 +47,7 @@
done
lemma acc_le_prodI [intro!]:
- "\<lbrakk> acc r\<^isub>A; acc r\<^isub>B \<rbrakk> \<Longrightarrow> acc(Product.le r\<^isub>A r\<^isub>B)"
+ "\<lbrakk> acc r\<^sub>A; acc r\<^sub>B \<rbrakk> \<Longrightarrow> acc(Product.le r\<^sub>A r\<^sub>B)"
apply (unfold acc_def)
apply (rule wf_subset)
apply (erule wf_lex_prod)
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Aug 13 16:25:47 2013 +0200
@@ -3352,7 +3352,7 @@
next
case False obtain y where "y\<in>s" and y:"\<forall>x\<in>s. dist z y \<le> dist z x"
using distance_attains_inf[OF assms(2) False] by auto
- show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) z + (norm(y - z))\<twosuperior> / 2" in exI)
+ show ?thesis apply(rule_tac x="y - z" in exI, rule_tac x="inner (y - z) z + (norm(y - z))\<^sup>2 / 2" in exI)
apply rule defer apply rule proof-
fix x assume "x\<in>s"
have "\<not> 0 < inner (z - y) (x - y)" apply(rule_tac notI) proof(drule closer_point_lemma)
@@ -3363,7 +3363,7 @@
using `x\<in>s` `y\<in>s` by (auto simp add: dist_commute algebra_simps) qed
moreover have "0 < norm (y - z) ^ 2" using `y\<in>s` `z\<notin>s` by auto
hence "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp
- ultimately show "inner (y - z) z + (norm (y - z))\<twosuperior> / 2 < inner (y - z) x"
+ ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x"
unfolding power2_norm_eq_inner and not_less by (auto simp add: field_simps inner_commute inner_diff)
qed(insert `y\<in>s` `z\<notin>s`, auto)
qed
--- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Aug 13 16:25:47 2013 +0200
@@ -879,7 +879,7 @@
obtains p where "p division_of {a..b}" "{c..d} \<in> p"
proof
let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a. {(\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) .. (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)}"
- def p \<equiv> "?B ` (Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)})"
+ def p \<equiv> "?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
show "{c .. d} \<in> p"
unfolding p_def
@@ -900,7 +900,7 @@
{
fix k
assume "k \<in> p"
- then obtain f where f: "f \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
+ then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
by (auto simp: p_def)
then show "\<exists>a b. k = {a..b}" by auto
have "k \<subseteq> {a..b} \<and> k \<noteq> {}"
@@ -917,7 +917,7 @@
then show "k \<noteq> {}" "k \<subseteq> {a .. b}" by auto
{
fix l assume "l \<in> p"
- then obtain g where g: "g \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
+ then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
by (auto simp: p_def)
assume "l \<noteq> k"
have "\<exists>i\<in>Basis. f i \<noteq> g i"
@@ -952,7 +952,7 @@
by auto
qed
then guess f unfolding bchoice_iff .. note f = this
- moreover then have "restrict f Basis \<in> Basis \<rightarrow>\<^isub>E {(a, c), (c, d), (d, b)}"
+ moreover then have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
by auto
moreover from f have "x \<in> ?B (restrict f Basis)"
by (auto simp: mem_interval eucl_le[where 'a='a])
--- a/src/HOL/Multivariate_Analysis/L2_Norm.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/L2_Norm.thy Tue Aug 13 16:25:47 2013 +0200
@@ -9,7 +9,7 @@
begin
definition
- "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
+ "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<^sup>2)"
lemma setL2_cong:
"\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
@@ -27,7 +27,7 @@
lemma setL2_insert [simp]:
"\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
- setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
+ setL2 f (insert a F) = sqrt ((f a)\<^sup>2 + (setL2 f F)\<^sup>2)"
unfolding setL2_def by (simp add: setsum_nonneg)
lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
@@ -94,12 +94,12 @@
show ?case by simp
next
case (insert x F)
- hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
- sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
+ hence "sqrt ((f x + g x)\<^sup>2 + (setL2 (\<lambda>i. f i + g i) F)\<^sup>2) \<le>
+ sqrt ((f x + g x)\<^sup>2 + (setL2 f F + setL2 g F)\<^sup>2)"
by (intro real_sqrt_le_mono add_left_mono power_mono insert
setL2_nonneg add_increasing zero_le_power2)
also have
- "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
+ "\<dots> \<le> sqrt ((f x)\<^sup>2 + (setL2 f F)\<^sup>2) + sqrt ((g x)\<^sup>2 + (setL2 g F)\<^sup>2)"
by (rule real_sqrt_sum_squares_triangle_ineq)
finally show ?case
using insert by simp
@@ -107,7 +107,7 @@
qed
lemma sqrt_sum_squares_le_sum:
- "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
+ "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<^sup>2 + y\<^sup>2) \<le> x + y"
apply (rule power2_le_imp_le)
apply (simp add: power2_sum mult_nonneg_nonneg)
apply simp
@@ -125,7 +125,7 @@
apply simp
done
-lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
+lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<^sup>2 + y\<^sup>2) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
apply (rule power2_le_imp_le)
apply (simp add: power2_sum mult_nonneg_nonneg)
apply simp
@@ -143,14 +143,14 @@
lemma setL2_mult_ineq_lemma:
fixes a b c d :: real
- shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+ shows "2 * (a * c) * (b * d) \<le> a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2"
proof -
- have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
- also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
+ have "0 \<le> (a * d - b * c)\<^sup>2" by simp
+ also have "\<dots> = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * d) * (b * c)"
by (simp only: power2_diff power_mult_distrib)
- also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
+ also have "\<dots> = a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2 - 2 * (a * c) * (b * d)"
by simp
- finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+ finally show "2 * (a * c) * (b * d) \<le> a\<^sup>2 * d\<^sup>2 + b\<^sup>2 * c\<^sup>2"
by simp
qed
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Aug 13 16:25:47 2013 +0200
@@ -94,11 +94,11 @@
lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
proof
assume "\<bar>x\<bar> \<le> \<bar>y\<bar>"
- then have "\<bar>x\<bar>\<twosuperior> \<le> \<bar>y\<bar>\<twosuperior>" by (rule power_mono, simp)
- then show "x\<twosuperior> \<le> y\<twosuperior>" by simp
+ then have "\<bar>x\<bar>\<^sup>2 \<le> \<bar>y\<bar>\<^sup>2" by (rule power_mono, simp)
+ then show "x\<^sup>2 \<le> y\<^sup>2" by simp
next
- assume "x\<twosuperior> \<le> y\<twosuperior>"
- then have "sqrt (x\<twosuperior>) \<le> sqrt (y\<twosuperior>)" by (rule real_sqrt_le_mono)
+ assume "x\<^sup>2 \<le> y\<^sup>2"
+ then have "sqrt (x\<^sup>2) \<le> sqrt (y\<^sup>2)" by (rule real_sqrt_le_mono)
then show "\<bar>x\<bar> \<le> \<bar>y\<bar>" by simp
qed
@@ -2601,7 +2601,7 @@
unfolding real_of_nat_def
apply(subst euclidean_inner)
apply (subst power2_abs[symmetric])
- apply (rule order_trans[OF setsum_bounded[where K="\<bar>infnorm x\<bar>\<twosuperior>"]])
+ apply (rule order_trans[OF setsum_bounded[where K="\<bar>infnorm x\<bar>\<^sup>2"]])
apply (auto simp add: power2_eq_square[symmetric])
apply (subst power2_abs[symmetric])
apply (rule power_mono)
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Aug 13 16:25:47 2013 +0200
@@ -745,7 +745,7 @@
show ?thesis
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)
fix y :: 'a assume *: "y \<in> box ?a ?b"
- have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<twosuperior>)"
+ have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)"
unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
@@ -756,9 +756,9 @@
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" by auto
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
unfolding e'_def by (auto simp: dist_real_def)
- then have "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
+ then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
by (rule power_strict_mono) auto
- then show "(dist (x \<bullet> i) (y \<bullet> i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
+ then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
by (simp add: power_divide)
qed auto
also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat)
@@ -771,7 +771,7 @@
assumes "open M"
defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)"
defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)"
- defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^isub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
+ defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"
proof -
{
@@ -4747,7 +4747,7 @@
proof-
obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"
using assms [unfolded bounded_def] by auto
- then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<twosuperior> + b\<twosuperior>)"
+ then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<^sup>2 + b\<^sup>2)"
by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)
thus ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto
qed
@@ -5500,7 +5500,7 @@
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" by simp
def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i"
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" by simp
- def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^isub>E (\<rat> \<times> \<rat>))"
+ def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))"
have "Ball B open" by (simp add: B_def open_box)
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))"
--- a/src/HOL/NSA/NSCA.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/NSA/NSCA.thy Tue Aug 13 16:25:47 2013 +0200
@@ -256,12 +256,12 @@
apply (erule (1) InfinitesimalD2)
done
-lemma real_sqrt_lessI: "\<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> sqrt x < u"
+lemma real_sqrt_lessI: "\<lbrakk>0 < u; x < u\<^sup>2\<rbrakk> \<Longrightarrow> sqrt x < u"
(* TODO: this belongs somewhere else *)
by (frule real_sqrt_less_mono) simp
lemma hypreal_sqrt_lessI:
- "\<And>x u. \<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u"
+ "\<And>x u. \<lbrakk>0 < u; x < u\<^sup>2\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u"
by transfer (rule real_sqrt_lessI)
lemma hypreal_sqrt_ge_zero: "\<And>x. 0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt) x"
@@ -270,7 +270,7 @@
lemma Infinitesimal_sqrt:
"\<lbrakk>x \<in> Infinitesimal; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal"
apply (rule InfinitesimalI2)
-apply (drule_tac r="r\<twosuperior>" in InfinitesimalD2, simp)
+apply (drule_tac r="r\<^sup>2" in InfinitesimalD2, simp)
apply (simp add: hypreal_sqrt_ge_zero)
apply (rule hypreal_sqrt_lessI, simp_all)
done
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Aug 13 16:25:47 2013 +0200
@@ -250,25 +250,25 @@
"loose (Lam t) k = loose t (Suc k)" |
"loose (App t u) k = (loose t k \<or> loose u k)"
-primrec subst\<^isub>1 where
-"subst\<^isub>1 \<sigma> (Var j) = \<sigma> j" |
-"subst\<^isub>1 \<sigma> (Lam t) =
- Lam (subst\<^isub>1 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 1) t)" |
-"subst\<^isub>1 \<sigma> (App t u) = App (subst\<^isub>1 \<sigma> t) (subst\<^isub>1 \<sigma> u)"
+primrec subst\<^sub>1 where
+"subst\<^sub>1 \<sigma> (Var j) = \<sigma> j" |
+"subst\<^sub>1 \<sigma> (Lam t) =
+ Lam (subst\<^sub>1 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 1) t)" |
+"subst\<^sub>1 \<sigma> (App t u) = App (subst\<^sub>1 \<sigma> t) (subst\<^sub>1 \<sigma> u)"
-lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>1 \<sigma> t = t"
+lemma "\<not> loose t 0 \<Longrightarrow> subst\<^sub>1 \<sigma> t = t"
nitpick [verbose, expect = genuine]
-nitpick [eval = "subst\<^isub>1 \<sigma> t", expect = genuine]
+nitpick [eval = "subst\<^sub>1 \<sigma> t", expect = genuine]
(* nitpick [dont_box, expect = unknown] *)
oops
-primrec subst\<^isub>2 where
-"subst\<^isub>2 \<sigma> (Var j) = \<sigma> j" |
-"subst\<^isub>2 \<sigma> (Lam t) =
- Lam (subst\<^isub>2 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 0) t)" |
-"subst\<^isub>2 \<sigma> (App t u) = App (subst\<^isub>2 \<sigma> t) (subst\<^isub>2 \<sigma> u)"
+primrec subst\<^sub>2 where
+"subst\<^sub>2 \<sigma> (Var j) = \<sigma> j" |
+"subst\<^sub>2 \<sigma> (Lam t) =
+ Lam (subst\<^sub>2 (\<lambda>n. case n of 0 \<Rightarrow> Var 0 | Suc m \<Rightarrow> lift (\<sigma> m) 0) t)" |
+"subst\<^sub>2 \<sigma> (App t u) = App (subst\<^sub>2 \<sigma> t) (subst\<^sub>2 \<sigma> u)"
-lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>2 \<sigma> t = t"
+lemma "\<not> loose t 0 \<Longrightarrow> subst\<^sub>2 \<sigma> t = t"
nitpick [card = 1\<emdash>5, expect = none]
sorry
@@ -354,73 +354,73 @@
datatype alphabet = a | b
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where
+ "[] \<in> S\<^sub>1"
+| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>1_sound:
+"w \<in> S\<^sub>1 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [expect = genuine]
oops
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where
+ "[] \<in> S\<^sub>2"
+| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2"
+| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2"
+| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
-theorem S\<^isub>2_sound:
-"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>2_sound:
+"w \<in> S\<^sub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [expect = genuine]
oops
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where
+ "[] \<in> S\<^sub>3"
+| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3"
+| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3"
+| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
-theorem S\<^isub>3_sound:
-"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>3_sound:
+"w \<in> S\<^sub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [card = 1\<emdash>5, expect = none]
sorry
-theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+theorem S\<^sub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>3"
nitpick [expect = genuine]
oops
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where
+ "[] \<in> S\<^sub>4"
+| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4"
+| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4"
+| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4"
+| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
-theorem S\<^isub>4_sound:
-"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>4_sound:
+"w \<in> S\<^sub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
nitpick [card = 1\<emdash>5, expect = none]
sorry
-theorem S\<^isub>4_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+theorem S\<^sub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>4"
nitpick [card = 1\<emdash>5, expect = none]
sorry
-theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete:
-"w \<in> S\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-"w \<in> A\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
-"w \<in> B\<^isub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
+theorem S\<^sub>4_A\<^sub>4_B\<^sub>4_sound_and_complete:
+"w \<in> S\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+"w \<in> A\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] + 1"
+"w \<in> B\<^sub>4 \<longleftrightarrow> length [x \<leftarrow> w. x = b] = length [x \<leftarrow> w. x = a] + 1"
nitpick [card = 1\<emdash>5, expect = none]
sorry
@@ -442,11 +442,11 @@
primrec left where
"left \<Lambda> = \<Lambda>" |
-"left (N _ _ t\<^isub>1 _) = t\<^isub>1"
+"left (N _ _ t\<^sub>1 _) = t\<^sub>1"
primrec right where
"right \<Lambda> = \<Lambda>" |
-"right (N _ _ _ t\<^isub>2) = t\<^isub>2"
+"right (N _ _ _ t\<^sub>2) = t\<^sub>2"
fun wf where
"wf \<Lambda> = True" |
@@ -484,31 +484,31 @@
nitpick [card = 1\<emdash>5, expect = none]
sorry
-primrec insort\<^isub>1 where
-"insort\<^isub>1 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" |
-"insort\<^isub>1 (N y k t u) x =
- (* (split \<circ> skew) *) (N y k (if x < y then insort\<^isub>1 t x else t)
- (if x > y then insort\<^isub>1 u x else u))"
+primrec insort\<^sub>1 where
+"insort\<^sub>1 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" |
+"insort\<^sub>1 (N y k t u) x =
+ (* (split \<circ> skew) *) (N y k (if x < y then insort\<^sub>1 t x else t)
+ (if x > y then insort\<^sub>1 u x else u))"
-theorem wf_insort\<^isub>1: "wf t \<Longrightarrow> wf (insort\<^isub>1 t x)"
+theorem wf_insort\<^sub>1: "wf t \<Longrightarrow> wf (insort\<^sub>1 t x)"
nitpick [expect = genuine]
oops
-theorem wf_insort\<^isub>1_nat: "wf t \<Longrightarrow> wf (insort\<^isub>1 t (x\<Colon>nat))"
-nitpick [eval = "insort\<^isub>1 t x", expect = genuine]
+theorem wf_insort\<^sub>1_nat: "wf t \<Longrightarrow> wf (insort\<^sub>1 t (x\<Colon>nat))"
+nitpick [eval = "insort\<^sub>1 t x", expect = genuine]
oops
-primrec insort\<^isub>2 where
-"insort\<^isub>2 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" |
-"insort\<^isub>2 (N y k t u) x =
- (split \<circ> skew) (N y k (if x < y then insort\<^isub>2 t x else t)
- (if x > y then insort\<^isub>2 u x else u))"
+primrec insort\<^sub>2 where
+"insort\<^sub>2 \<Lambda> x = N x 1 \<Lambda> \<Lambda>" |
+"insort\<^sub>2 (N y k t u) x =
+ (split \<circ> skew) (N y k (if x < y then insort\<^sub>2 t x else t)
+ (if x > y then insort\<^sub>2 u x else u))"
-theorem wf_insort\<^isub>2: "wf t \<Longrightarrow> wf (insort\<^isub>2 t x)"
+theorem wf_insort\<^sub>2: "wf t \<Longrightarrow> wf (insort\<^sub>2 t x)"
nitpick [card = 1\<emdash>5, expect = none]
sorry
-theorem dataset_insort\<^isub>2: "dataset (insort\<^isub>2 t x) = {x} \<union> dataset t"
+theorem dataset_insort\<^sub>2: "dataset (insort\<^sub>2 t x) = {x} \<union> dataset t"
nitpick [card = 1\<emdash>5, expect = none]
sorry
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy Tue Aug 13 16:25:47 2013 +0200
@@ -131,8 +131,8 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\<Colon>nat).
- b\<^isub>1 < b\<^isub>11 \<and> f5 g x = f5 (\<lambda>a. if b\<^isub>1 < b\<^isub>11 then a else h b\<^isub>2) x"
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ b\<^sub>1 < b\<^sub>11 \<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 < b\<^sub>11 then a else h b\<^sub>2) x"
nitpick [expect = potential]
nitpick [dont_specialize, expect = none]
nitpick [dont_box, expect = none]
@@ -140,13 +140,13 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\<Colon>nat).
- b\<^isub>1 < b\<^isub>11
- \<and> f5 g x = f5 (\<lambda>a. if b\<^isub>1 < b\<^isub>11 then
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ b\<^sub>1 < b\<^sub>11
+ \<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 < b\<^sub>11 then
a
else
- h b\<^isub>2 + h b\<^isub>3 + h b\<^isub>4 + h b\<^isub>5 + h b\<^isub>6 + h b\<^isub>7 + h b\<^isub>8
- + h b\<^isub>9 + h b\<^isub>10) x"
+ h b\<^sub>2 + h b\<^sub>3 + h b\<^sub>4 + h b\<^sub>5 + h b\<^sub>6 + h b\<^sub>7 + h b\<^sub>8
+ + h b\<^sub>9 + h b\<^sub>10) x"
nitpick [card nat = 2, card 'a = 1, expect = none]
nitpick [card nat = 2, card 'a = 1, dont_box, expect = none]
nitpick [card nat = 2, card 'a = 1, dont_specialize, expect = none]
@@ -154,13 +154,13 @@
sorry
lemma "\<forall>a. g a = a
- \<Longrightarrow> \<exists>b\<^isub>1 b\<^isub>2 b\<^isub>3 b\<^isub>4 b\<^isub>5 b\<^isub>6 b\<^isub>7 b\<^isub>8 b\<^isub>9 b\<^isub>10 (b\<^isub>11\<Colon>nat).
- b\<^isub>1 < b\<^isub>11
- \<and> f5 g x = f5 (\<lambda>a. if b\<^isub>1 \<ge> b\<^isub>11 then
+ \<Longrightarrow> \<exists>b\<^sub>1 b\<^sub>2 b\<^sub>3 b\<^sub>4 b\<^sub>5 b\<^sub>6 b\<^sub>7 b\<^sub>8 b\<^sub>9 b\<^sub>10 (b\<^sub>11\<Colon>nat).
+ b\<^sub>1 < b\<^sub>11
+ \<and> f5 g x = f5 (\<lambda>a. if b\<^sub>1 \<ge> b\<^sub>11 then
a
else
- h b\<^isub>2 + h b\<^isub>3 + h b\<^isub>4 + h b\<^isub>5 + h b\<^isub>6 + h b\<^isub>7 + h b\<^isub>8
- + h b\<^isub>9 + h b\<^isub>10) x"
+ h b\<^sub>2 + h b\<^sub>3 + h b\<^sub>4 + h b\<^sub>5 + h b\<^sub>6 + h b\<^sub>7 + h b\<^sub>8
+ + h b\<^sub>9 + h b\<^sub>10) x"
nitpick [card nat = 2, card 'a = 1, expect = potential]
nitpick [card nat = 2, card 'a = 1, dont_box, expect = potential]
nitpick [card nat = 2, card 'a = 1, dont_specialize, expect = potential]
--- a/src/HOL/Nominal/Examples/CK_Machine.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/CK_Machine.thy Tue Aug 13 16:25:47 2013 +0200
@@ -43,11 +43,11 @@
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
where
"(VAR x)[y::=s] = (if x=y then s else (VAR x))"
-| "(APP t\<^isub>1 t\<^isub>2)[y::=s] = APP (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
+| "(APP t\<^sub>1 t\<^sub>2)[y::=s] = APP (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])"
| "x\<sharp>(y,s) \<Longrightarrow> (LAM [x].t)[y::=s] = LAM [x].(t[y::=s])"
| "(NUM n)[y::=s] = NUM n"
-| "(t\<^isub>1 -- t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) -- (t\<^isub>2[y::=s])"
-| "(t\<^isub>1 ++ t\<^isub>2)[y::=s] = (t\<^isub>1[y::=s]) ++ (t\<^isub>2[y::=s])"
+| "(t\<^sub>1 -- t\<^sub>2)[y::=s] = (t\<^sub>1[y::=s]) -- (t\<^sub>2[y::=s])"
+| "(t\<^sub>1 ++ t\<^sub>2)[y::=s] = (t\<^sub>1[y::=s]) ++ (t\<^sub>2[y::=s])"
| "x\<sharp>(y,s) \<Longrightarrow> (FIX [x].t)[y::=s] = FIX [x].(t[y::=s])"
| "TRUE[y::=s] = TRUE"
| "FALSE[y::=s] = FALSE"
@@ -387,7 +387,7 @@
abbreviation
"sub_tctx" :: "tctx \<Rightarrow> tctx \<Rightarrow> bool" ("_ \<subseteq> _")
where
- "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^isub>1 \<longrightarrow> x \<in> set \<Gamma>\<^isub>2"
+ "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
text {* Valid Typing Contexts *}
@@ -429,11 +429,11 @@
typing :: "tctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
where
t_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> VAR x : T"
-| t_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> APP t\<^isub>1 t\<^isub>2 : T\<^isub>2"
-| t_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> LAM [x].t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| t_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> APP t\<^sub>1 t\<^sub>2 : T\<^sub>2"
+| t_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> LAM [x].t : T\<^sub>1 \<rightarrow> T\<^sub>2"
| t_NUM[intro]: "\<Gamma> \<turnstile> (NUM n) : tINT"
-| t_DIFF[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : tINT; \<Gamma> \<turnstile> t\<^isub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 -- t\<^isub>2 : tINT"
-| t_PLUS[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : tINT; \<Gamma> \<turnstile> t\<^isub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 ++ t\<^isub>2 : tINT"
+| t_DIFF[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : tINT; \<Gamma> \<turnstile> t\<^sub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 -- t\<^sub>2 : tINT"
+| t_PLUS[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : tINT; \<Gamma> \<turnstile> t\<^sub>2 : tINT\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 ++ t\<^sub>2 : tINT"
| t_TRUE[intro]: "\<Gamma> \<turnstile> TRUE : tBOOL"
| t_FALSE[intro]: "\<Gamma> \<turnstile> FALSE : tBOOL"
| t_IF[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : tBOOL; \<Gamma> \<turnstile> t2 : T; \<Gamma> \<turnstile> t3 : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> IF t1 t2 t3 : T"
@@ -443,8 +443,8 @@
declare lam.inject[simp]
inductive_cases typing_inversion[elim]:
- "\<Gamma> \<turnstile> t\<^isub>1 -- t\<^isub>2 : T"
- "\<Gamma> \<turnstile> t\<^isub>1 ++ t\<^isub>2 : T"
+ "\<Gamma> \<turnstile> t\<^sub>1 -- t\<^sub>2 : T"
+ "\<Gamma> \<turnstile> t\<^sub>1 ++ t\<^sub>2 : T"
"\<Gamma> \<turnstile> IF t1 t2 t3 : T"
"\<Gamma> \<turnstile> ZET t : T"
"\<Gamma> \<turnstile> EQI t1 t2 : T"
@@ -461,7 +461,7 @@
lemma t_LAM_inversion[dest]:
assumes ty: "\<Gamma> \<turnstile> LAM [x].t : T"
and fc: "x\<sharp>\<Gamma>"
- shows "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \<rightarrow> T\<^isub>2 \<and> (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2"
+ shows "\<exists>T\<^sub>1 T\<^sub>2. T = T\<^sub>1 \<rightarrow> T\<^sub>2 \<and> (x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2"
using ty fc
by (cases rule: typing.strong_cases)
(auto simp add: alpha lam.inject abs_fresh ty_fresh)
--- a/src/HOL/Nominal/Examples/CR.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/CR.thy Tue Aug 13 16:25:47 2013 +0200
@@ -146,12 +146,12 @@
section {* Beta Reduction *}
inductive
- "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
+ "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
where
- b1[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)"
- | b2[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)"
- | b3[intro]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [a].s2)"
- | b4[intro]: "a\<sharp>s2 \<Longrightarrow> (App (Lam [a].s1) s2)\<longrightarrow>\<^isub>\<beta>(s1[a::=s2])"
+ b1[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^sub>\<beta>(App s2 t)"
+ | b2[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^sub>\<beta>(App t s2)"
+ | b3[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^sub>\<beta> (Lam [a].s2)"
+ | b4[intro]: "a\<sharp>s2 \<Longrightarrow> (App (Lam [a].s1) s2)\<longrightarrow>\<^sub>\<beta>(s1[a::=s2])"
equivariance Beta
@@ -159,29 +159,29 @@
by (simp_all add: abs_fresh fresh_fact')
inductive
- "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80)
+ "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
where
- bs1[intro, simp]: "M \<longrightarrow>\<^isub>\<beta>\<^sup>* M"
- | bs2[intro]: "\<lbrakk>M1\<longrightarrow>\<^isub>\<beta>\<^sup>* M2; M2 \<longrightarrow>\<^isub>\<beta> M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
+ bs1[intro, simp]: "M \<longrightarrow>\<^sub>\<beta>\<^sup>* M"
+ | bs2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>\<beta>\<^sup>* M2; M2 \<longrightarrow>\<^sub>\<beta> M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M3"
equivariance Beta_star
lemma beta_star_trans:
- assumes a1: "M1\<longrightarrow>\<^isub>\<beta>\<^sup>* M2"
- and a2: "M2\<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
- shows "M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M3"
+ assumes a1: "M1\<longrightarrow>\<^sub>\<beta>\<^sup>* M2"
+ and a2: "M2\<longrightarrow>\<^sub>\<beta>\<^sup>* M3"
+ shows "M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M3"
using a2 a1
by (induct) (auto)
section {* One-Reduction *}
inductive
- One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1 _" [80,80] 80)
+ One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
where
- o1[intro!]: "M\<longrightarrow>\<^isub>1M"
- | o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^isub>1t2;s1\<longrightarrow>\<^isub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^isub>1(App t2 s2)"
- | o3[simp,intro!]: "s1\<longrightarrow>\<^isub>1s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>1(Lam [a].s2)"
- | o4[simp,intro!]: "\<lbrakk>a\<sharp>(s1,s2); s1\<longrightarrow>\<^isub>1s2;t1\<longrightarrow>\<^isub>1t2\<rbrakk> \<Longrightarrow> (App (Lam [a].t1) s1)\<longrightarrow>\<^isub>1(t2[a::=s2])"
+ o1[intro!]: "M\<longrightarrow>\<^sub>1M"
+ | o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2;s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^sub>1(App t2 s2)"
+ | o3[simp,intro!]: "s1\<longrightarrow>\<^sub>1s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^sub>1(Lam [a].s2)"
+ | o4[simp,intro!]: "\<lbrakk>a\<sharp>(s1,s2); s1\<longrightarrow>\<^sub>1s2;t1\<longrightarrow>\<^sub>1t2\<rbrakk> \<Longrightarrow> (App (Lam [a].t1) s1)\<longrightarrow>\<^sub>1(t2[a::=s2])"
equivariance One
@@ -189,23 +189,23 @@
by (simp_all add: abs_fresh fresh_fact')
inductive
- "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1\<^sup>* _" [80,80] 80)
+ "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
where
- os1[intro, simp]: "M \<longrightarrow>\<^isub>1\<^sup>* M"
- | os2[intro]: "\<lbrakk>M1\<longrightarrow>\<^isub>1\<^sup>* M2; M2 \<longrightarrow>\<^isub>1 M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^isub>1\<^sup>* M3"
+ os1[intro, simp]: "M \<longrightarrow>\<^sub>1\<^sup>* M"
+ | os2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>1\<^sup>* M2; M2 \<longrightarrow>\<^sub>1 M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>1\<^sup>* M3"
equivariance One_star
lemma one_star_trans:
- assumes a1: "M1\<longrightarrow>\<^isub>1\<^sup>* M2"
- and a2: "M2\<longrightarrow>\<^isub>1\<^sup>* M3"
- shows "M1\<longrightarrow>\<^isub>1\<^sup>* M3"
+ assumes a1: "M1\<longrightarrow>\<^sub>1\<^sup>* M2"
+ and a2: "M2\<longrightarrow>\<^sub>1\<^sup>* M3"
+ shows "M1\<longrightarrow>\<^sub>1\<^sup>* M3"
using a2 a1
by (induct) (auto)
lemma one_fresh_preserv:
fixes a :: "name"
- assumes a: "t\<longrightarrow>\<^isub>1s"
+ assumes a: "t\<longrightarrow>\<^sub>1s"
and b: "a\<sharp>t"
shows "a\<sharp>s"
using a b
@@ -247,7 +247,7 @@
lemma one_fresh_preserv_automatic:
fixes a :: "name"
- assumes a: "t\<longrightarrow>\<^isub>1s"
+ assumes a: "t\<longrightarrow>\<^sub>1s"
and b: "a\<sharp>t"
shows "a\<sharp>s"
using a b
@@ -263,8 +263,8 @@
(auto simp add: calc_atm fresh_atm abs_fresh)
lemma one_abs:
- assumes a: "Lam [a].t\<longrightarrow>\<^isub>1t'"
- shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>1t''"
+ assumes a: "Lam [a].t\<longrightarrow>\<^sub>1t'"
+ shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^sub>1t''"
proof -
have "a\<sharp>Lam [a].t" by (simp add: abs_fresh)
with a have "a\<sharp>t'" by (simp add: one_fresh_preserv)
@@ -274,15 +274,15 @@
qed
lemma one_app:
- assumes a: "App t1 t2 \<longrightarrow>\<^isub>1 t'"
- shows "(\<exists>s1 s2. t' = App s1 s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or>
- (\<exists>a s s1 s2. t1 = Lam [a].s \<and> t' = s1[a::=s2] \<and> s \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2 \<and> a\<sharp>(t2,s2))"
+ assumes a: "App t1 t2 \<longrightarrow>\<^sub>1 t'"
+ shows "(\<exists>s1 s2. t' = App s1 s2 \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2) \<or>
+ (\<exists>a s s1 s2. t1 = Lam [a].s \<and> t' = s1[a::=s2] \<and> s \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2 \<and> a\<sharp>(t2,s2))"
using a by (erule_tac One.cases) (auto simp add: lam.inject)
lemma one_red:
- assumes a: "App (Lam [a].t1) t2 \<longrightarrow>\<^isub>1 M" "a\<sharp>(t2,M)"
- shows "(\<exists>s1 s2. M = App (Lam [a].s1) s2 \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2) \<or>
- (\<exists>s1 s2. M = s1[a::=s2] \<and> t1 \<longrightarrow>\<^isub>1 s1 \<and> t2 \<longrightarrow>\<^isub>1 s2)"
+ assumes a: "App (Lam [a].t1) t2 \<longrightarrow>\<^sub>1 M" "a\<sharp>(t2,M)"
+ shows "(\<exists>s1 s2. M = App (Lam [a].s1) s2 \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2) \<or>
+ (\<exists>s1 s2. M = s1[a::=s2] \<and> t1 \<longrightarrow>\<^sub>1 s1 \<and> t2 \<longrightarrow>\<^sub>1 s2)"
using a
by (cases rule: One.strong_cases [where a="a" and aa="a"])
(auto dest: one_abs simp add: lam.inject abs_fresh alpha fresh_prod)
@@ -290,31 +290,31 @@
text {* first case in Lemma 3.2.4*}
lemma one_subst_aux:
- assumes a: "N\<longrightarrow>\<^isub>1N'"
- shows "M[x::=N] \<longrightarrow>\<^isub>1 M[x::=N']"
+ assumes a: "N\<longrightarrow>\<^sub>1N'"
+ shows "M[x::=N] \<longrightarrow>\<^sub>1 M[x::=N']"
using a
proof (nominal_induct M avoiding: x N N' rule: lam.strong_induct)
case (Var y)
- thus "Var y[x::=N] \<longrightarrow>\<^isub>1 Var y[x::=N']" by (cases "x=y") auto
+ thus "Var y[x::=N] \<longrightarrow>\<^sub>1 Var y[x::=N']" by (cases "x=y") auto
next
case (App P Q) (* application case - third line *)
- thus "(App P Q)[x::=N] \<longrightarrow>\<^isub>1 (App P Q)[x::=N']" using o2 by simp
+ thus "(App P Q)[x::=N] \<longrightarrow>\<^sub>1 (App P Q)[x::=N']" using o2 by simp
next
case (Lam y P) (* abstraction case - fourth line *)
- thus "(Lam [y].P)[x::=N] \<longrightarrow>\<^isub>1 (Lam [y].P)[x::=N']" using o3 by simp
+ thus "(Lam [y].P)[x::=N] \<longrightarrow>\<^sub>1 (Lam [y].P)[x::=N']" using o3 by simp
qed
lemma one_subst_aux_automatic:
- assumes a: "N\<longrightarrow>\<^isub>1N'"
- shows "M[x::=N] \<longrightarrow>\<^isub>1 M[x::=N']"
+ assumes a: "N\<longrightarrow>\<^sub>1N'"
+ shows "M[x::=N] \<longrightarrow>\<^sub>1 M[x::=N']"
using a
by (nominal_induct M avoiding: x N N' rule: lam.strong_induct)
(auto simp add: fresh_prod fresh_atm)
lemma one_subst:
- assumes a: "M\<longrightarrow>\<^isub>1M'"
- and b: "N\<longrightarrow>\<^isub>1N'"
- shows "M[x::=N]\<longrightarrow>\<^isub>1M'[x::=N']"
+ assumes a: "M\<longrightarrow>\<^sub>1M'"
+ and b: "N\<longrightarrow>\<^sub>1N'"
+ shows "M[x::=N]\<longrightarrow>\<^sub>1M'[x::=N']"
using a b
proof (nominal_induct M M' avoiding: N N' x rule: One.strong_induct)
case (o1 M)
@@ -328,22 +328,22 @@
next
case (o4 a N1 N2 M1 M2 N N' x)
have vc: "a\<sharp>N" "a\<sharp>N'" "a\<sharp>x" "a\<sharp>N1" "a\<sharp>N2" by fact+
- have asm: "N\<longrightarrow>\<^isub>1N'" by fact
+ have asm: "N\<longrightarrow>\<^sub>1N'" by fact
show ?case
proof -
have "(App (Lam [a].M1) N1)[x::=N] = App (Lam [a].(M1[x::=N])) (N1[x::=N])" using vc by simp
- moreover have "App (Lam [a].(M1[x::=N])) (N1[x::=N]) \<longrightarrow>\<^isub>1 M2[x::=N'][a::=N2[x::=N']]"
+ moreover have "App (Lam [a].(M1[x::=N])) (N1[x::=N]) \<longrightarrow>\<^sub>1 M2[x::=N'][a::=N2[x::=N']]"
using o4 asm by (simp add: fresh_fact)
moreover have "M2[x::=N'][a::=N2[x::=N']] = M2[a::=N2][x::=N']"
using vc by (simp add: substitution_lemma fresh_atm)
- ultimately show "(App (Lam [a].M1) N1)[x::=N] \<longrightarrow>\<^isub>1 M2[a::=N2][x::=N']" by simp
+ ultimately show "(App (Lam [a].M1) N1)[x::=N] \<longrightarrow>\<^sub>1 M2[a::=N2][x::=N']" by simp
qed
qed
lemma one_subst_automatic:
- assumes a: "M\<longrightarrow>\<^isub>1M'"
- and b: "N\<longrightarrow>\<^isub>1N'"
- shows "M[x::=N]\<longrightarrow>\<^isub>1M'[x::=N']"
+ assumes a: "M\<longrightarrow>\<^sub>1M'"
+ and b: "N\<longrightarrow>\<^sub>1N'"
+ shows "M[x::=N]\<longrightarrow>\<^sub>1M'[x::=N']"
using a b
by (nominal_induct M M' avoiding: N N' x rule: One.strong_induct)
(auto simp add: one_subst_aux substitution_lemma fresh_atm fresh_fact)
@@ -351,122 +351,122 @@
lemma diamond[rule_format]:
fixes M :: "lam"
and M1:: "lam"
- assumes a: "M\<longrightarrow>\<^isub>1M1"
- and b: "M\<longrightarrow>\<^isub>1M2"
- shows "\<exists>M3. M1\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3"
+ assumes a: "M\<longrightarrow>\<^sub>1M1"
+ and b: "M\<longrightarrow>\<^sub>1M2"
+ shows "\<exists>M3. M1\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3"
using a b
proof (nominal_induct avoiding: M1 M2 rule: One.strong_induct)
case (o1 M) (* case 1 --- M1 = M *)
- thus "\<exists>M3. M\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" by blast
+ thus "\<exists>M3. M\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast
next
case (o4 x Q Q' P P') (* case 2 --- a beta-reduction occurs*)
have vc: "x\<sharp>Q" "x\<sharp>Q'" "x\<sharp>M2" by fact+
- have i1: "\<And>M2. Q \<longrightarrow>\<^isub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)" by fact
- have i2: "\<And>M2. P \<longrightarrow>\<^isub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)" by fact
- have "App (Lam [x].P) Q \<longrightarrow>\<^isub>1 M2" by fact
- hence "(\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q') \<or>
- (\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q')" using vc by (simp add: one_red)
+ have i1: "\<And>M2. Q \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact
+ have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact
+ have "App (Lam [x].P) Q \<longrightarrow>\<^sub>1 M2" by fact
+ hence "(\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q') \<or>
+ (\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q')" using vc by (simp add: one_red)
moreover (* subcase 2.1 *)
- { assume "\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q'"
+ { assume "\<exists>P' Q'. M2 = App (Lam [x].P') Q' \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q'"
then obtain P'' and Q'' where
- b1: "M2=App (Lam [x].P'') Q''" and b2: "P\<longrightarrow>\<^isub>1P''" and b3: "Q\<longrightarrow>\<^isub>1Q''" by blast
- from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
+ b1: "M2=App (Lam [x].P'') Q''" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast
+ from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp
then obtain P''' where
- c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by force
- from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
+ c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by force
+ from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp
then obtain Q''' where
- d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by force
+ d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by force
from c1 c2 d1 d2
- have "P'[x::=Q']\<longrightarrow>\<^isub>1P'''[x::=Q'''] \<and> App (Lam [x].P'') Q'' \<longrightarrow>\<^isub>1 P'''[x::=Q''']"
+ have "P'[x::=Q']\<longrightarrow>\<^sub>1P'''[x::=Q'''] \<and> App (Lam [x].P'') Q'' \<longrightarrow>\<^sub>1 P'''[x::=Q''']"
using vc b3 by (auto simp add: one_subst one_fresh_preserv)
- hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by blast
+ hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast
}
moreover (* subcase 2.2 *)
- { assume "\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^isub>1P' \<and> Q\<longrightarrow>\<^isub>1Q'"
+ { assume "\<exists>P' Q'. M2 = P'[x::=Q'] \<and> P\<longrightarrow>\<^sub>1P' \<and> Q\<longrightarrow>\<^sub>1Q'"
then obtain P'' Q'' where
- b1: "M2=P''[x::=Q'']" and b2: "P\<longrightarrow>\<^isub>1P''" and b3: "Q\<longrightarrow>\<^isub>1Q''" by blast
- from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
+ b1: "M2=P''[x::=Q'']" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast
+ from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp
then obtain P''' where
- c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by blast
- from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
+ c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by blast
+ from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp
then obtain Q''' where
- d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by blast
+ d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast
from c1 c2 d1 d2
- have "P'[x::=Q']\<longrightarrow>\<^isub>1P'''[x::=Q'''] \<and> P''[x::=Q'']\<longrightarrow>\<^isub>1P'''[x::=Q''']"
+ have "P'[x::=Q']\<longrightarrow>\<^sub>1P'''[x::=Q'''] \<and> P''[x::=Q'']\<longrightarrow>\<^sub>1P'''[x::=Q''']"
by (force simp add: one_subst)
- hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by blast
+ hence "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast
}
- ultimately show "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" by blast
+ ultimately show "\<exists>M3. P'[x::=Q']\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast
next
case (o2 P P' Q Q') (* case 3 *)
- have i0: "P\<longrightarrow>\<^isub>1P'" by fact
- have i0': "Q\<longrightarrow>\<^isub>1Q'" by fact
- have i1: "\<And>M2. Q \<longrightarrow>\<^isub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)" by fact
- have i2: "\<And>M2. P \<longrightarrow>\<^isub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)" by fact
- assume "App P Q \<longrightarrow>\<^isub>1 M2"
- hence "(\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^isub>1P'' \<and> Q\<longrightarrow>\<^isub>1Q'') \<or>
- (\<exists>x P' P'' Q'. P = Lam [x].P' \<and> M2 = P''[x::=Q'] \<and> P'\<longrightarrow>\<^isub>1 P'' \<and> Q\<longrightarrow>\<^isub>1Q' \<and> x\<sharp>(Q,Q'))"
+ have i0: "P\<longrightarrow>\<^sub>1P'" by fact
+ have i0': "Q\<longrightarrow>\<^sub>1Q'" by fact
+ have i1: "\<And>M2. Q \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact
+ have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact
+ assume "App P Q \<longrightarrow>\<^sub>1 M2"
+ hence "(\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^sub>1P'' \<and> Q\<longrightarrow>\<^sub>1Q'') \<or>
+ (\<exists>x P' P'' Q'. P = Lam [x].P' \<and> M2 = P''[x::=Q'] \<and> P'\<longrightarrow>\<^sub>1 P'' \<and> Q\<longrightarrow>\<^sub>1Q' \<and> x\<sharp>(Q,Q'))"
by (simp add: one_app[simplified])
moreover (* subcase 3.1 *)
- { assume "\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^isub>1P'' \<and> Q\<longrightarrow>\<^isub>1Q''"
+ { assume "\<exists>P'' Q''. M2 = App P'' Q'' \<and> P\<longrightarrow>\<^sub>1P'' \<and> Q\<longrightarrow>\<^sub>1Q''"
then obtain P'' and Q'' where
- b1: "M2=App P'' Q''" and b2: "P\<longrightarrow>\<^isub>1P''" and b3: "Q\<longrightarrow>\<^isub>1Q''" by blast
- from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> P''\<longrightarrow>\<^isub>1M3)" by simp
+ b1: "M2=App P'' Q''" and b2: "P\<longrightarrow>\<^sub>1P''" and b3: "Q\<longrightarrow>\<^sub>1Q''" by blast
+ from b2 i2 have "(\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> P''\<longrightarrow>\<^sub>1M3)" by simp
then obtain P''' where
- c1: "P'\<longrightarrow>\<^isub>1P'''" and c2: "P''\<longrightarrow>\<^isub>1P'''" by blast
- from b3 i1 have "\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3" by simp
+ c1: "P'\<longrightarrow>\<^sub>1P'''" and c2: "P''\<longrightarrow>\<^sub>1P'''" by blast
+ from b3 i1 have "\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3" by simp
then obtain Q''' where
- d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by blast
+ d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast
from c1 c2 d1 d2
- have "App P' Q'\<longrightarrow>\<^isub>1App P''' Q''' \<and> App P'' Q'' \<longrightarrow>\<^isub>1 App P''' Q'''" by blast
- hence "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by blast
+ have "App P' Q'\<longrightarrow>\<^sub>1App P''' Q''' \<and> App P'' Q'' \<longrightarrow>\<^sub>1 App P''' Q'''" by blast
+ hence "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast
}
moreover (* subcase 3.2 *)
- { assume "\<exists>x P1 P'' Q''. P = Lam [x].P1 \<and> M2 = P''[x::=Q''] \<and> P1\<longrightarrow>\<^isub>1 P'' \<and> Q\<longrightarrow>\<^isub>1Q'' \<and> x\<sharp>(Q,Q'')"
+ { assume "\<exists>x P1 P'' Q''. P = Lam [x].P1 \<and> M2 = P''[x::=Q''] \<and> P1\<longrightarrow>\<^sub>1 P'' \<and> Q\<longrightarrow>\<^sub>1Q'' \<and> x\<sharp>(Q,Q'')"
then obtain x P1 P1'' Q'' where
b0: "P = Lam [x].P1" and b1: "M2 = P1''[x::=Q'']" and
- b2: "P1\<longrightarrow>\<^isub>1P1''" and b3: "Q\<longrightarrow>\<^isub>1Q''" and vc: "x\<sharp>(Q,Q'')" by blast
- from b0 i0 have "\<exists>P1'. P'=Lam [x].P1' \<and> P1\<longrightarrow>\<^isub>1P1'" by (simp add: one_abs)
- then obtain P1' where g1: "P'=Lam [x].P1'" and g2: "P1\<longrightarrow>\<^isub>1P1'" by blast
- from g1 b0 b2 i2 have "(\<exists>M3. (Lam [x].P1')\<longrightarrow>\<^isub>1M3 \<and> (Lam [x].P1'')\<longrightarrow>\<^isub>1M3)" by simp
+ b2: "P1\<longrightarrow>\<^sub>1P1''" and b3: "Q\<longrightarrow>\<^sub>1Q''" and vc: "x\<sharp>(Q,Q'')" by blast
+ from b0 i0 have "\<exists>P1'. P'=Lam [x].P1' \<and> P1\<longrightarrow>\<^sub>1P1'" by (simp add: one_abs)
+ then obtain P1' where g1: "P'=Lam [x].P1'" and g2: "P1\<longrightarrow>\<^sub>1P1'" by blast
+ from g1 b0 b2 i2 have "(\<exists>M3. (Lam [x].P1')\<longrightarrow>\<^sub>1M3 \<and> (Lam [x].P1'')\<longrightarrow>\<^sub>1M3)" by simp
then obtain P1''' where
- c1: "(Lam [x].P1')\<longrightarrow>\<^isub>1P1'''" and c2: "(Lam [x].P1'')\<longrightarrow>\<^isub>1P1'''" by blast
- from c1 have "\<exists>R1. P1'''=Lam [x].R1 \<and> P1'\<longrightarrow>\<^isub>1R1" by (simp add: one_abs)
- then obtain R1 where r1: "P1'''=Lam [x].R1" and r2: "P1'\<longrightarrow>\<^isub>1R1" by blast
- from c2 have "\<exists>R2. P1'''=Lam [x].R2 \<and> P1''\<longrightarrow>\<^isub>1R2" by (simp add: one_abs)
- then obtain R2 where r3: "P1'''=Lam [x].R2" and r4: "P1''\<longrightarrow>\<^isub>1R2" by blast
+ c1: "(Lam [x].P1')\<longrightarrow>\<^sub>1P1'''" and c2: "(Lam [x].P1'')\<longrightarrow>\<^sub>1P1'''" by blast
+ from c1 have "\<exists>R1. P1'''=Lam [x].R1 \<and> P1'\<longrightarrow>\<^sub>1R1" by (simp add: one_abs)
+ then obtain R1 where r1: "P1'''=Lam [x].R1" and r2: "P1'\<longrightarrow>\<^sub>1R1" by blast
+ from c2 have "\<exists>R2. P1'''=Lam [x].R2 \<and> P1''\<longrightarrow>\<^sub>1R2" by (simp add: one_abs)
+ then obtain R2 where r3: "P1'''=Lam [x].R2" and r4: "P1''\<longrightarrow>\<^sub>1R2" by blast
from r1 r3 have r5: "R1=R2" by (simp add: lam.inject alpha)
- from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^isub>1M3 \<and> Q''\<longrightarrow>\<^isub>1M3)" by simp
+ from b3 i1 have "(\<exists>M3. Q'\<longrightarrow>\<^sub>1M3 \<and> Q''\<longrightarrow>\<^sub>1M3)" by simp
then obtain Q''' where
- d1: "Q'\<longrightarrow>\<^isub>1Q'''" and d2: "Q''\<longrightarrow>\<^isub>1Q'''" by blast
+ d1: "Q'\<longrightarrow>\<^sub>1Q'''" and d2: "Q''\<longrightarrow>\<^sub>1Q'''" by blast
from g1 r2 d1 r4 r5 d2
- have "App P' Q'\<longrightarrow>\<^isub>1R1[x::=Q'''] \<and> P1''[x::=Q'']\<longrightarrow>\<^isub>1R1[x::=Q''']"
+ have "App P' Q'\<longrightarrow>\<^sub>1R1[x::=Q'''] \<and> P1''[x::=Q'']\<longrightarrow>\<^sub>1R1[x::=Q''']"
using vc i0' by (simp add: one_subst one_fresh_preserv)
- hence "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 by blast
+ hence "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 by blast
}
- ultimately show "\<exists>M3. App P' Q'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" by blast
+ ultimately show "\<exists>M3. App P' Q'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" by blast
next
case (o3 P P' x) (* case 4 *)
- have i1: "P\<longrightarrow>\<^isub>1P'" by fact
- have i2: "\<And>M2. P \<longrightarrow>\<^isub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3)" by fact
- have "(Lam [x].P)\<longrightarrow>\<^isub>1 M2" by fact
- hence "\<exists>P''. M2=Lam [x].P'' \<and> P\<longrightarrow>\<^isub>1P''" by (simp add: one_abs)
- then obtain P'' where b1: "M2=Lam [x].P''" and b2: "P\<longrightarrow>\<^isub>1P''" by blast
- from i2 b1 b2 have "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^isub>1M3 \<and> (Lam [x].P'')\<longrightarrow>\<^isub>1M3" by blast
- then obtain M3 where c1: "(Lam [x].P')\<longrightarrow>\<^isub>1M3" and c2: "(Lam [x].P'')\<longrightarrow>\<^isub>1M3" by blast
- from c1 have "\<exists>R1. M3=Lam [x].R1 \<and> P'\<longrightarrow>\<^isub>1R1" by (simp add: one_abs)
- then obtain R1 where r1: "M3=Lam [x].R1" and r2: "P'\<longrightarrow>\<^isub>1R1" by blast
- from c2 have "\<exists>R2. M3=Lam [x].R2 \<and> P''\<longrightarrow>\<^isub>1R2" by (simp add: one_abs)
- then obtain R2 where r3: "M3=Lam [x].R2" and r4: "P''\<longrightarrow>\<^isub>1R2" by blast
+ have i1: "P\<longrightarrow>\<^sub>1P'" by fact
+ have i2: "\<And>M2. P \<longrightarrow>\<^sub>1M2 \<Longrightarrow> (\<exists>M3. P'\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3)" by fact
+ have "(Lam [x].P)\<longrightarrow>\<^sub>1 M2" by fact
+ hence "\<exists>P''. M2=Lam [x].P'' \<and> P\<longrightarrow>\<^sub>1P''" by (simp add: one_abs)
+ then obtain P'' where b1: "M2=Lam [x].P''" and b2: "P\<longrightarrow>\<^sub>1P''" by blast
+ from i2 b1 b2 have "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^sub>1M3 \<and> (Lam [x].P'')\<longrightarrow>\<^sub>1M3" by blast
+ then obtain M3 where c1: "(Lam [x].P')\<longrightarrow>\<^sub>1M3" and c2: "(Lam [x].P'')\<longrightarrow>\<^sub>1M3" by blast
+ from c1 have "\<exists>R1. M3=Lam [x].R1 \<and> P'\<longrightarrow>\<^sub>1R1" by (simp add: one_abs)
+ then obtain R1 where r1: "M3=Lam [x].R1" and r2: "P'\<longrightarrow>\<^sub>1R1" by blast
+ from c2 have "\<exists>R2. M3=Lam [x].R2 \<and> P''\<longrightarrow>\<^sub>1R2" by (simp add: one_abs)
+ then obtain R2 where r3: "M3=Lam [x].R2" and r4: "P''\<longrightarrow>\<^sub>1R2" by blast
from r1 r3 have r5: "R1=R2" by (simp add: lam.inject alpha)
- from r2 r4 have "(Lam [x].P')\<longrightarrow>\<^isub>1(Lam [x].R1) \<and> (Lam [x].P'')\<longrightarrow>\<^isub>1(Lam [x].R2)"
+ from r2 r4 have "(Lam [x].P')\<longrightarrow>\<^sub>1(Lam [x].R1) \<and> (Lam [x].P'')\<longrightarrow>\<^sub>1(Lam [x].R2)"
by (simp add: one_subst)
- thus "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^isub>1M3 \<and> M2\<longrightarrow>\<^isub>1M3" using b1 r5 by blast
+ thus "\<exists>M3. (Lam [x].P')\<longrightarrow>\<^sub>1M3 \<and> M2\<longrightarrow>\<^sub>1M3" using b1 r5 by blast
qed
lemma one_lam_cong:
- assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
- shows "(Lam [a].t1)\<longrightarrow>\<^isub>\<beta>\<^sup>*(Lam [a].t2)"
+ assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2"
+ shows "(Lam [a].t1)\<longrightarrow>\<^sub>\<beta>\<^sup>*(Lam [a].t2)"
using a
proof induct
case bs1 thus ?case by simp
@@ -476,8 +476,8 @@
qed
lemma one_app_congL:
- assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
- shows "App t1 s\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s"
+ assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2"
+ shows "App t1 s\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s"
using a
proof induct
case bs1 thus ?case by simp
@@ -486,8 +486,8 @@
qed
lemma one_app_congR:
- assumes a: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
- shows "App s t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App s t2"
+ assumes a: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2"
+ shows "App s t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App s t2"
using a
proof induct
case bs1 thus ?case by simp
@@ -496,19 +496,19 @@
qed
lemma one_app_cong:
- assumes a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
- and a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2"
- shows "App t1 s1\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2"
+ assumes a1: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2"
+ and a2: "s1\<longrightarrow>\<^sub>\<beta>\<^sup>*s2"
+ shows "App t1 s1\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2"
proof -
- have "App t1 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s1" using a1 by (rule one_app_congL)
+ have "App t1 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s1" using a1 by (rule one_app_congL)
moreover
- have "App t2 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2" using a2 by (rule one_app_congR)
+ have "App t2 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2" using a2 by (rule one_app_congR)
ultimately show ?thesis by (rule beta_star_trans)
qed
lemma one_beta_star:
- assumes a: "(t1\<longrightarrow>\<^isub>1t2)"
- shows "(t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2)"
+ assumes a: "(t1\<longrightarrow>\<^sub>1t2)"
+ shows "(t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2)"
using a
proof(nominal_induct rule: One.strong_induct)
case o1 thus ?case by simp
@@ -519,16 +519,16 @@
next
case (o4 a s1 s2 t1 t2)
have vc: "a\<sharp>s1" "a\<sharp>s2" by fact+
- have a1: "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t2" and a2: "s1\<longrightarrow>\<^isub>\<beta>\<^sup>*s2" by fact+
- have c1: "(App (Lam [a].t2) s2) \<longrightarrow>\<^isub>\<beta> (t2 [a::= s2])" using vc by (simp add: b4)
- from a1 a2 have c2: "App (Lam [a].t1 ) s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App (Lam [a].t2 ) s2"
+ have a1: "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t2" and a2: "s1\<longrightarrow>\<^sub>\<beta>\<^sup>*s2" by fact+
+ have c1: "(App (Lam [a].t2) s2) \<longrightarrow>\<^sub>\<beta> (t2 [a::= s2])" using vc by (simp add: b4)
+ from a1 a2 have c2: "App (Lam [a].t1 ) s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App (Lam [a].t2 ) s2"
by (blast intro!: one_app_cong one_lam_cong)
show ?case using c2 c1 by (blast intro: beta_star_trans)
qed
lemma one_star_lam_cong:
- assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
- shows "(Lam [a].t1)\<longrightarrow>\<^isub>1\<^sup>* (Lam [a].t2)"
+ assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2"
+ shows "(Lam [a].t1)\<longrightarrow>\<^sub>1\<^sup>* (Lam [a].t2)"
using a
proof induct
case os1 thus ?case by simp
@@ -537,8 +537,8 @@
qed
lemma one_star_app_congL:
- assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
- shows "App t1 s\<longrightarrow>\<^isub>1\<^sup>* App t2 s"
+ assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2"
+ shows "App t1 s\<longrightarrow>\<^sub>1\<^sup>* App t2 s"
using a
proof induct
case os1 thus ?case by simp
@@ -547,8 +547,8 @@
qed
lemma one_star_app_congR:
- assumes a: "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
- shows "App s t1 \<longrightarrow>\<^isub>1\<^sup>* App s t2"
+ assumes a: "t1\<longrightarrow>\<^sub>1\<^sup>*t2"
+ shows "App s t1 \<longrightarrow>\<^sub>1\<^sup>* App s t2"
using a
proof induct
case os1 thus ?case by simp
@@ -557,8 +557,8 @@
qed
lemma beta_one_star:
- assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2"
- shows "t1\<longrightarrow>\<^isub>1\<^sup>*t2"
+ assumes a: "t1\<longrightarrow>\<^sub>\<beta>t2"
+ shows "t1\<longrightarrow>\<^sub>1\<^sup>*t2"
using a
proof(induct)
case b1 thus ?case by (blast intro!: one_star_app_congL)
@@ -571,88 +571,88 @@
qed
lemma trans_closure:
- shows "(M1\<longrightarrow>\<^isub>1\<^sup>*M2) = (M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2)"
+ shows "(M1\<longrightarrow>\<^sub>1\<^sup>*M2) = (M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2)"
proof
- assume "M1 \<longrightarrow>\<^isub>1\<^sup>* M2"
- then show "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2"
+ assume "M1 \<longrightarrow>\<^sub>1\<^sup>* M2"
+ then show "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2"
proof induct
- case (os1 M1) thus "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M1" by simp
+ case (os1 M1) thus "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M1" by simp
next
case (os2 M1 M2 M3)
- have "M2\<longrightarrow>\<^isub>1M3" by fact
- then have "M2\<longrightarrow>\<^isub>\<beta>\<^sup>*M3" by (rule one_beta_star)
- moreover have "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M2" by fact
- ultimately show "M1\<longrightarrow>\<^isub>\<beta>\<^sup>*M3" by (auto intro: beta_star_trans)
+ have "M2\<longrightarrow>\<^sub>1M3" by fact
+ then have "M2\<longrightarrow>\<^sub>\<beta>\<^sup>*M3" by (rule one_beta_star)
+ moreover have "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M2" by fact
+ ultimately show "M1\<longrightarrow>\<^sub>\<beta>\<^sup>*M3" by (auto intro: beta_star_trans)
qed
next
- assume "M1 \<longrightarrow>\<^isub>\<beta>\<^sup>* M2"
- then show "M1\<longrightarrow>\<^isub>1\<^sup>*M2"
+ assume "M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M2"
+ then show "M1\<longrightarrow>\<^sub>1\<^sup>*M2"
proof induct
- case (bs1 M1) thus "M1\<longrightarrow>\<^isub>1\<^sup>*M1" by simp
+ case (bs1 M1) thus "M1\<longrightarrow>\<^sub>1\<^sup>*M1" by simp
next
case (bs2 M1 M2 M3)
- have "M2\<longrightarrow>\<^isub>\<beta>M3" by fact
- then have "M2\<longrightarrow>\<^isub>1\<^sup>*M3" by (rule beta_one_star)
- moreover have "M1\<longrightarrow>\<^isub>1\<^sup>*M2" by fact
- ultimately show "M1\<longrightarrow>\<^isub>1\<^sup>*M3" by (auto intro: one_star_trans)
+ have "M2\<longrightarrow>\<^sub>\<beta>M3" by fact
+ then have "M2\<longrightarrow>\<^sub>1\<^sup>*M3" by (rule beta_one_star)
+ moreover have "M1\<longrightarrow>\<^sub>1\<^sup>*M2" by fact
+ ultimately show "M1\<longrightarrow>\<^sub>1\<^sup>*M3" by (auto intro: one_star_trans)
qed
qed
lemma cr_one:
- assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t1"
- and b: "t\<longrightarrow>\<^isub>1t2"
- shows "\<exists>t3. t1\<longrightarrow>\<^isub>1t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3"
+ assumes a: "t\<longrightarrow>\<^sub>1\<^sup>*t1"
+ and b: "t\<longrightarrow>\<^sub>1t2"
+ shows "\<exists>t3. t1\<longrightarrow>\<^sub>1t3 \<and> t2\<longrightarrow>\<^sub>1\<^sup>*t3"
using a b
proof (induct arbitrary: t2)
case os1 thus ?case by force
next
case (os2 t s1 s2 t2)
- have b: "s1 \<longrightarrow>\<^isub>1 s2" by fact
- have h: "\<And>t2. t \<longrightarrow>\<^isub>1 t2 \<Longrightarrow> (\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3)" by fact
- have c: "t \<longrightarrow>\<^isub>1 t2" by fact
- show "\<exists>t3. s2 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3"
+ have b: "s1 \<longrightarrow>\<^sub>1 s2" by fact
+ have h: "\<And>t2. t \<longrightarrow>\<^sub>1 t2 \<Longrightarrow> (\<exists>t3. s1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3)" by fact
+ have c: "t \<longrightarrow>\<^sub>1 t2" by fact
+ show "\<exists>t3. s2 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3"
proof -
- from c h have "\<exists>t3. s1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by blast
- then obtain t3 where c1: "s1 \<longrightarrow>\<^isub>1 t3" and c2: "t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by blast
- have "\<exists>t4. s2 \<longrightarrow>\<^isub>1 t4 \<and> t3 \<longrightarrow>\<^isub>1 t4" using b c1 by (blast intro: diamond)
+ from c h have "\<exists>t3. s1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by blast
+ then obtain t3 where c1: "s1 \<longrightarrow>\<^sub>1 t3" and c2: "t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by blast
+ have "\<exists>t4. s2 \<longrightarrow>\<^sub>1 t4 \<and> t3 \<longrightarrow>\<^sub>1 t4" using b c1 by (blast intro: diamond)
thus ?thesis using c2 by (blast intro: one_star_trans)
qed
qed
lemma cr_one_star:
- assumes a: "t\<longrightarrow>\<^isub>1\<^sup>*t2"
- and b: "t\<longrightarrow>\<^isub>1\<^sup>*t1"
- shows "\<exists>t3. t1\<longrightarrow>\<^isub>1\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>1\<^sup>*t3"
+ assumes a: "t\<longrightarrow>\<^sub>1\<^sup>*t2"
+ and b: "t\<longrightarrow>\<^sub>1\<^sup>*t1"
+ shows "\<exists>t3. t1\<longrightarrow>\<^sub>1\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>1\<^sup>*t3"
using a b
proof (induct arbitrary: t1)
case (os1 t) then show ?case by force
next
case (os2 t s1 s2 t1)
- have c: "t \<longrightarrow>\<^isub>1\<^sup>* s1" by fact
- have c': "t \<longrightarrow>\<^isub>1\<^sup>* t1" by fact
- have d: "s1 \<longrightarrow>\<^isub>1 s2" by fact
- have "t \<longrightarrow>\<^isub>1\<^sup>* t1 \<Longrightarrow> (\<exists>t3. t1 \<longrightarrow>\<^isub>1\<^sup>* t3 \<and> s1 \<longrightarrow>\<^isub>1\<^sup>* t3)" by fact
- then obtain t3 where f1: "t1 \<longrightarrow>\<^isub>1\<^sup>* t3"
- and f2: "s1 \<longrightarrow>\<^isub>1\<^sup>* t3" using c' by blast
- from cr_one d f2 have "\<exists>t4. t3\<longrightarrow>\<^isub>1t4 \<and> s2\<longrightarrow>\<^isub>1\<^sup>*t4" by blast
- then obtain t4 where g1: "t3\<longrightarrow>\<^isub>1t4"
- and g2: "s2\<longrightarrow>\<^isub>1\<^sup>*t4" by blast
- have "t1\<longrightarrow>\<^isub>1\<^sup>*t4" using f1 g1 by (blast intro: one_star_trans)
+ have c: "t \<longrightarrow>\<^sub>1\<^sup>* s1" by fact
+ have c': "t \<longrightarrow>\<^sub>1\<^sup>* t1" by fact
+ have d: "s1 \<longrightarrow>\<^sub>1 s2" by fact
+ have "t \<longrightarrow>\<^sub>1\<^sup>* t1 \<Longrightarrow> (\<exists>t3. t1 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> s1 \<longrightarrow>\<^sub>1\<^sup>* t3)" by fact
+ then obtain t3 where f1: "t1 \<longrightarrow>\<^sub>1\<^sup>* t3"
+ and f2: "s1 \<longrightarrow>\<^sub>1\<^sup>* t3" using c' by blast
+ from cr_one d f2 have "\<exists>t4. t3\<longrightarrow>\<^sub>1t4 \<and> s2\<longrightarrow>\<^sub>1\<^sup>*t4" by blast
+ then obtain t4 where g1: "t3\<longrightarrow>\<^sub>1t4"
+ and g2: "s2\<longrightarrow>\<^sub>1\<^sup>*t4" by blast
+ have "t1\<longrightarrow>\<^sub>1\<^sup>*t4" using f1 g1 by (blast intro: one_star_trans)
thus ?case using g2 by blast
qed
lemma cr_beta_star:
- assumes a1: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t1"
- and a2: "t\<longrightarrow>\<^isub>\<beta>\<^sup>*t2"
- shows "\<exists>t3. t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3"
+ assumes a1: "t\<longrightarrow>\<^sub>\<beta>\<^sup>*t1"
+ and a2: "t\<longrightarrow>\<^sub>\<beta>\<^sup>*t2"
+ shows "\<exists>t3. t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3"
proof -
- from a1 have "t\<longrightarrow>\<^isub>1\<^sup>*t1" by (simp only: trans_closure)
+ from a1 have "t\<longrightarrow>\<^sub>1\<^sup>*t1" by (simp only: trans_closure)
moreover
- from a2 have "t\<longrightarrow>\<^isub>1\<^sup>*t2" by (simp only: trans_closure)
- ultimately have "\<exists>t3. t1\<longrightarrow>\<^isub>1\<^sup>*t3 \<and> t2\<longrightarrow>\<^isub>1\<^sup>*t3" by (blast intro: cr_one_star)
- then obtain t3 where "t1\<longrightarrow>\<^isub>1\<^sup>*t3" and "t2\<longrightarrow>\<^isub>1\<^sup>*t3" by blast
- hence "t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" and "t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by (simp_all only: trans_closure)
- then show "\<exists>t3. t1\<longrightarrow>\<^isub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^isub>\<beta>\<^sup>*t3" by blast
+ from a2 have "t\<longrightarrow>\<^sub>1\<^sup>*t2" by (simp only: trans_closure)
+ ultimately have "\<exists>t3. t1\<longrightarrow>\<^sub>1\<^sup>*t3 \<and> t2\<longrightarrow>\<^sub>1\<^sup>*t3" by (blast intro: cr_one_star)
+ then obtain t3 where "t1\<longrightarrow>\<^sub>1\<^sup>*t3" and "t2\<longrightarrow>\<^sub>1\<^sup>*t3" by blast
+ hence "t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" and "t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" by (simp_all only: trans_closure)
+ then show "\<exists>t3. t1\<longrightarrow>\<^sub>\<beta>\<^sup>*t3\<and>t2\<longrightarrow>\<^sub>\<beta>\<^sup>*t3" by blast
qed
end
--- a/src/HOL/Nominal/Examples/CR_Takahashi.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/CR_Takahashi.thy Tue Aug 13 16:25:47 2013 +0200
@@ -22,7 +22,7 @@
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
where
"(Var x)[y::=s] = (if x=y then s else (Var x))"
-| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
+| "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])"
| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
apply(finite_guess)+
apply(rule TrueI)+
@@ -66,81 +66,81 @@
section {* Beta-Reduction *}
inductive
- "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
+ "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
where
- b1[intro]: "t1 \<longrightarrow>\<^isub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^isub>\<beta> App t2 s"
-| b2[intro]: "s1 \<longrightarrow>\<^isub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^isub>\<beta> App t s2"
-| b3[intro]: "t1 \<longrightarrow>\<^isub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^isub>\<beta> Lam [x].t2"
-| b4[intro]: "App (Lam [x].t) s \<longrightarrow>\<^isub>\<beta> t[x::=s]"
+ b1[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^sub>\<beta> App t2 s"
+| b2[intro]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
+| b3[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>\<beta> Lam [x].t2"
+| b4[intro]: "App (Lam [x].t) s \<longrightarrow>\<^sub>\<beta> t[x::=s]"
section {* Transitive Closure of Beta *}
inductive
- "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80)
+ "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
where
- bs1[intro]: "t \<longrightarrow>\<^isub>\<beta>\<^sup>* t"
-| bs2[intro]: "t \<longrightarrow>\<^isub>\<beta> s \<Longrightarrow> t \<longrightarrow>\<^isub>\<beta>\<^sup>* s"
-| bs3[intro,trans]: "\<lbrakk>t1\<longrightarrow>\<^isub>\<beta>\<^sup>* t2; t2 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3"
+ bs1[intro]: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t"
+| bs2[intro]: "t \<longrightarrow>\<^sub>\<beta> s \<Longrightarrow> t \<longrightarrow>\<^sub>\<beta>\<^sup>* s"
+| bs3[intro,trans]: "\<lbrakk>t1\<longrightarrow>\<^sub>\<beta>\<^sup>* t2; t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3"
section {* One-Reduction *}
inductive
- One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1 _" [80,80] 80)
+ One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
where
- o1[intro]: "Var x\<longrightarrow>\<^isub>1 Var x"
-| o2[intro]: "\<lbrakk>t1\<longrightarrow>\<^isub>1t2; s1\<longrightarrow>\<^isub>1s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^isub>1 App t2 s2"
-| o3[intro]: "t1\<longrightarrow>\<^isub>1t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^isub>1 Lam [x].t2"
-| o4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1\<longrightarrow>\<^isub>1t2; s1\<longrightarrow>\<^isub>1s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^isub>1 t2[x::=s2]"
+ o1[intro]: "Var x\<longrightarrow>\<^sub>1 Var x"
+| o2[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2; s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^sub>1 App t2 s2"
+| o3[intro]: "t1\<longrightarrow>\<^sub>1t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>1 Lam [x].t2"
+| o4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1\<longrightarrow>\<^sub>1t2; s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]"
equivariance One
nominal_inductive One
by (simp_all add: abs_fresh fresh_fact)
lemma One_refl:
- shows "t \<longrightarrow>\<^isub>1 t"
+ shows "t \<longrightarrow>\<^sub>1 t"
by (nominal_induct t rule: lam.strong_induct) (auto)
lemma One_subst:
- assumes a: "t1 \<longrightarrow>\<^isub>1 t2" "s1 \<longrightarrow>\<^isub>1 s2"
- shows "t1[x::=s1] \<longrightarrow>\<^isub>1 t2[x::=s2]"
+ assumes a: "t1 \<longrightarrow>\<^sub>1 t2" "s1 \<longrightarrow>\<^sub>1 s2"
+ shows "t1[x::=s1] \<longrightarrow>\<^sub>1 t2[x::=s2]"
using a
by (nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct)
(auto simp add: substitution_lemma fresh_atm fresh_fact)
lemma better_o4_intro:
- assumes a: "t1 \<longrightarrow>\<^isub>1 t2" "s1 \<longrightarrow>\<^isub>1 s2"
- shows "App (Lam [x].t1) s1 \<longrightarrow>\<^isub>1 t2[x::=s2]"
+ assumes a: "t1 \<longrightarrow>\<^sub>1 t2" "s1 \<longrightarrow>\<^sub>1 s2"
+ shows "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]"
proof -
obtain y::"name" where fs: "y\<sharp>(x,t1,s1,t2,s2)" by (rule exists_fresh, rule fin_supp, blast)
have "App (Lam [x].t1) s1 = App (Lam [y].([(y,x)]\<bullet>t1)) s1" using fs
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>1 ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: One.eqvt)
+ also have "\<dots> \<longrightarrow>\<^sub>1 ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: One.eqvt)
also have "\<dots> = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric])
- finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^isub>1 t2[x::=s2]" by simp
+ finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]" by simp
qed
lemma One_Var:
- assumes a: "Var x \<longrightarrow>\<^isub>1 M"
+ assumes a: "Var x \<longrightarrow>\<^sub>1 M"
shows "M = Var x"
using a by (cases rule: One.cases) (simp_all)
lemma One_Lam:
- assumes a: "Lam [x].t \<longrightarrow>\<^isub>1 s" "x\<sharp>s"
- shows "\<exists>t'. s = Lam [x].t' \<and> t \<longrightarrow>\<^isub>1 t'"
+ assumes a: "Lam [x].t \<longrightarrow>\<^sub>1 s" "x\<sharp>s"
+ shows "\<exists>t'. s = Lam [x].t' \<and> t \<longrightarrow>\<^sub>1 t'"
using a
by (cases rule: One.strong_cases)
(auto simp add: lam.inject abs_fresh alpha)
lemma One_App:
- assumes a: "App t s \<longrightarrow>\<^isub>1 r"
- shows "(\<exists>t' s'. r = App t' s' \<and> t \<longrightarrow>\<^isub>1 t' \<and> s \<longrightarrow>\<^isub>1 s') \<or>
- (\<exists>x p p' s'. r = p'[x::=s'] \<and> t = Lam [x].p \<and> p \<longrightarrow>\<^isub>1 p' \<and> s \<longrightarrow>\<^isub>1 s' \<and> x\<sharp>(s,s'))"
+ assumes a: "App t s \<longrightarrow>\<^sub>1 r"
+ shows "(\<exists>t' s'. r = App t' s' \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s') \<or>
+ (\<exists>x p p' s'. r = p'[x::=s'] \<and> t = Lam [x].p \<and> p \<longrightarrow>\<^sub>1 p' \<and> s \<longrightarrow>\<^sub>1 s' \<and> x\<sharp>(s,s'))"
using a by (cases rule: One.cases) (auto simp add: lam.inject)
lemma One_Redex:
- assumes a: "App (Lam [x].t) s \<longrightarrow>\<^isub>1 r" "x\<sharp>(s,r)"
- shows "(\<exists>t' s'. r = App (Lam [x].t') s' \<and> t \<longrightarrow>\<^isub>1 t' \<and> s \<longrightarrow>\<^isub>1 s') \<or>
- (\<exists>t' s'. r = t'[x::=s'] \<and> t \<longrightarrow>\<^isub>1 t' \<and> s \<longrightarrow>\<^isub>1 s')"
+ assumes a: "App (Lam [x].t) s \<longrightarrow>\<^sub>1 r" "x\<sharp>(s,r)"
+ shows "(\<exists>t' s'. r = App (Lam [x].t') s' \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s') \<or>
+ (\<exists>t' s'. r = t'[x::=s'] \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s')"
using a
by (cases rule: One.strong_cases)
(auto dest: One_Lam simp add: lam.inject abs_fresh alpha fresh_prod)
@@ -148,162 +148,162 @@
section {* Transitive Closure of One *}
inductive
- "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>1\<^sup>* _" [80,80] 80)
+ "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
where
- os1[intro]: "t \<longrightarrow>\<^isub>1\<^sup>* t"
-| os2[intro]: "t \<longrightarrow>\<^isub>1 s \<Longrightarrow> t \<longrightarrow>\<^isub>1\<^sup>* s"
-| os3[intro]: "\<lbrakk>t1\<longrightarrow>\<^isub>1\<^sup>* t2; t2 \<longrightarrow>\<^isub>1\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^isub>1\<^sup>* t3"
+ os1[intro]: "t \<longrightarrow>\<^sub>1\<^sup>* t"
+| os2[intro]: "t \<longrightarrow>\<^sub>1 s \<Longrightarrow> t \<longrightarrow>\<^sub>1\<^sup>* s"
+| os3[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1\<^sup>* t2; t2 \<longrightarrow>\<^sub>1\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>1\<^sup>* t3"
section {* Complete Development Reduction *}
inductive
- Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>\<^isub>d _" [80,80]80)
+ Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>\<^sub>d _" [80,80]80)
where
- d1[intro]: "Var x \<longrightarrow>\<^isub>d Var x"
-| d2[intro]: "t \<longrightarrow>\<^isub>d s \<Longrightarrow> Lam [x].t \<longrightarrow>\<^isub>d Lam[x].s"
-| d3[intro]: "\<lbrakk>\<not>(\<exists>y t'. t1 = Lam [y].t'); t1 \<longrightarrow>\<^isub>d t2; s1 \<longrightarrow>\<^isub>d s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^isub>d App t2 s2"
-| d4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1 \<longrightarrow>\<^isub>d t2; s1 \<longrightarrow>\<^isub>d s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^isub>d t2[x::=s2]"
+ d1[intro]: "Var x \<longrightarrow>\<^sub>d Var x"
+| d2[intro]: "t \<longrightarrow>\<^sub>d s \<Longrightarrow> Lam [x].t \<longrightarrow>\<^sub>d Lam[x].s"
+| d3[intro]: "\<lbrakk>\<not>(\<exists>y t'. t1 = Lam [y].t'); t1 \<longrightarrow>\<^sub>d t2; s1 \<longrightarrow>\<^sub>d s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^sub>d App t2 s2"
+| d4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1 \<longrightarrow>\<^sub>d t2; s1 \<longrightarrow>\<^sub>d s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]"
equivariance Dev
nominal_inductive Dev
by (simp_all add: abs_fresh fresh_fact)
lemma better_d4_intro:
- assumes a: "t1 \<longrightarrow>\<^isub>d t2" "s1 \<longrightarrow>\<^isub>d s2"
- shows "App (Lam [x].t1) s1 \<longrightarrow>\<^isub>d t2[x::=s2]"
+ assumes a: "t1 \<longrightarrow>\<^sub>d t2" "s1 \<longrightarrow>\<^sub>d s2"
+ shows "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]"
proof -
obtain y::"name" where fs: "y\<sharp>(x,t1,s1,t2,s2)" by (rule exists_fresh, rule fin_supp,blast)
have "App (Lam [x].t1) s1 = App (Lam [y].([(y,x)]\<bullet>t1)) s1" using fs
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>d ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: Dev.eqvt)
+ also have "\<dots> \<longrightarrow>\<^sub>d ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: Dev.eqvt)
also have "\<dots> = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric])
- finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^isub>d t2[x::=s2]" by simp
+ finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]" by simp
qed
lemma Dev_preserves_fresh:
fixes x::"name"
- assumes a: "M\<longrightarrow>\<^isub>d N"
+ assumes a: "M\<longrightarrow>\<^sub>d N"
shows "x\<sharp>M \<Longrightarrow> x\<sharp>N"
using a
by (induct) (auto simp add: abs_fresh fresh_fact)
lemma Dev_Lam:
- assumes a: "Lam [x].M \<longrightarrow>\<^isub>d N"
- shows "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^isub>d N'"
+ assumes a: "Lam [x].M \<longrightarrow>\<^sub>d N"
+ shows "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^sub>d N'"
proof -
from a have "x\<sharp>Lam [x].M" by (simp add: abs_fresh)
with a have "x\<sharp>N" by (simp add: Dev_preserves_fresh)
- with a show "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^isub>d N'"
+ with a show "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^sub>d N'"
by (cases rule: Dev.strong_cases)
(auto simp add: lam.inject abs_fresh alpha)
qed
lemma Development_existence:
- shows "\<exists>M'. M \<longrightarrow>\<^isub>d M'"
+ shows "\<exists>M'. M \<longrightarrow>\<^sub>d M'"
by (nominal_induct M rule: lam.strong_induct)
(auto dest!: Dev_Lam intro: better_d4_intro)
lemma Triangle:
- assumes a: "t \<longrightarrow>\<^isub>d t1" "t \<longrightarrow>\<^isub>1 t2"
- shows "t2 \<longrightarrow>\<^isub>1 t1"
+ assumes a: "t \<longrightarrow>\<^sub>d t1" "t \<longrightarrow>\<^sub>1 t2"
+ shows "t2 \<longrightarrow>\<^sub>1 t1"
using a
proof(nominal_induct avoiding: t2 rule: Dev.strong_induct)
case (d4 x s1 s2 t1 t1' t2)
have fc: "x\<sharp>t2" "x\<sharp>s1" by fact+
- have "App (Lam [x].t1) s1 \<longrightarrow>\<^isub>1 t2" by fact
+ have "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2" by fact
then obtain t' s' where reds:
- "(t2 = App (Lam [x].t') s' \<and> t1 \<longrightarrow>\<^isub>1 t' \<and> s1 \<longrightarrow>\<^isub>1 s') \<or>
- (t2 = t'[x::=s'] \<and> t1 \<longrightarrow>\<^isub>1 t' \<and> s1 \<longrightarrow>\<^isub>1 s')"
+ "(t2 = App (Lam [x].t') s' \<and> t1 \<longrightarrow>\<^sub>1 t' \<and> s1 \<longrightarrow>\<^sub>1 s') \<or>
+ (t2 = t'[x::=s'] \<and> t1 \<longrightarrow>\<^sub>1 t' \<and> s1 \<longrightarrow>\<^sub>1 s')"
using fc by (auto dest!: One_Redex)
- have ih1: "t1 \<longrightarrow>\<^isub>1 t' \<Longrightarrow> t' \<longrightarrow>\<^isub>1 t1'" by fact
- have ih2: "s1 \<longrightarrow>\<^isub>1 s' \<Longrightarrow> s' \<longrightarrow>\<^isub>1 s2" by fact
- { assume "t1 \<longrightarrow>\<^isub>1 t'" "s1 \<longrightarrow>\<^isub>1 s'"
- then have "App (Lam [x].t') s' \<longrightarrow>\<^isub>1 t1'[x::=s2]"
+ have ih1: "t1 \<longrightarrow>\<^sub>1 t' \<Longrightarrow> t' \<longrightarrow>\<^sub>1 t1'" by fact
+ have ih2: "s1 \<longrightarrow>\<^sub>1 s' \<Longrightarrow> s' \<longrightarrow>\<^sub>1 s2" by fact
+ { assume "t1 \<longrightarrow>\<^sub>1 t'" "s1 \<longrightarrow>\<^sub>1 s'"
+ then have "App (Lam [x].t') s' \<longrightarrow>\<^sub>1 t1'[x::=s2]"
using ih1 ih2 by (auto intro: better_o4_intro)
}
moreover
- { assume "t1 \<longrightarrow>\<^isub>1 t'" "s1 \<longrightarrow>\<^isub>1 s'"
- then have "t'[x::=s'] \<longrightarrow>\<^isub>1 t1'[x::=s2]"
+ { assume "t1 \<longrightarrow>\<^sub>1 t'" "s1 \<longrightarrow>\<^sub>1 s'"
+ then have "t'[x::=s'] \<longrightarrow>\<^sub>1 t1'[x::=s2]"
using ih1 ih2 by (auto intro: One_subst)
}
- ultimately show "t2 \<longrightarrow>\<^isub>1 t1'[x::=s2]" using reds by auto
+ ultimately show "t2 \<longrightarrow>\<^sub>1 t1'[x::=s2]" using reds by auto
qed (auto dest!: One_Lam One_Var One_App)
lemma Diamond_for_One:
- assumes a: "t \<longrightarrow>\<^isub>1 t1" "t \<longrightarrow>\<^isub>1 t2"
- shows "\<exists>t3. t2 \<longrightarrow>\<^isub>1 t3 \<and> t1 \<longrightarrow>\<^isub>1 t3"
+ assumes a: "t \<longrightarrow>\<^sub>1 t1" "t \<longrightarrow>\<^sub>1 t2"
+ shows "\<exists>t3. t2 \<longrightarrow>\<^sub>1 t3 \<and> t1 \<longrightarrow>\<^sub>1 t3"
proof -
- obtain tc where "t \<longrightarrow>\<^isub>d tc" using Development_existence by blast
- with a have "t2 \<longrightarrow>\<^isub>1 tc" and "t1 \<longrightarrow>\<^isub>1 tc" by (simp_all add: Triangle)
- then show "\<exists>t3. t2 \<longrightarrow>\<^isub>1 t3 \<and> t1 \<longrightarrow>\<^isub>1 t3" by blast
+ obtain tc where "t \<longrightarrow>\<^sub>d tc" using Development_existence by blast
+ with a have "t2 \<longrightarrow>\<^sub>1 tc" and "t1 \<longrightarrow>\<^sub>1 tc" by (simp_all add: Triangle)
+ then show "\<exists>t3. t2 \<longrightarrow>\<^sub>1 t3 \<and> t1 \<longrightarrow>\<^sub>1 t3" by blast
qed
lemma Rectangle_for_One:
- assumes a: "t \<longrightarrow>\<^isub>1\<^sup>* t1" "t \<longrightarrow>\<^isub>1 t2"
- shows "\<exists>t3. t1 \<longrightarrow>\<^isub>1 t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3"
+ assumes a: "t \<longrightarrow>\<^sub>1\<^sup>* t1" "t \<longrightarrow>\<^sub>1 t2"
+ shows "\<exists>t3. t1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3"
using a Diamond_for_One by (induct arbitrary: t2) (blast)+
lemma CR_for_One_star:
- assumes a: "t \<longrightarrow>\<^isub>1\<^sup>* t1" "t \<longrightarrow>\<^isub>1\<^sup>* t2"
- shows "\<exists>t3. t2 \<longrightarrow>\<^isub>1\<^sup>* t3 \<and> t1 \<longrightarrow>\<^isub>1\<^sup>* t3"
+ assumes a: "t \<longrightarrow>\<^sub>1\<^sup>* t1" "t \<longrightarrow>\<^sub>1\<^sup>* t2"
+ shows "\<exists>t3. t2 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> t1 \<longrightarrow>\<^sub>1\<^sup>* t3"
using a Rectangle_for_One by (induct arbitrary: t2) (blast)+
section {* Establishing the Equivalence of Beta-star and One-star *}
lemma Beta_Lam_cong:
- assumes a: "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2"
- shows "Lam [x].t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* Lam [x].t2"
+ assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
+ shows "Lam [x].t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* Lam [x].t2"
using a by (induct) (blast)+
lemma Beta_App_cong_aux:
- assumes a: "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2"
- shows "App t1 s\<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s"
- and "App s t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App s t2"
+ assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
+ shows "App t1 s\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s"
+ and "App s t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App s t2"
using a by (induct) (blast)+
lemma Beta_App_cong:
- assumes a: "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2" "s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* s2"
- shows "App t1 s1 \<longrightarrow>\<^isub>\<beta>\<^sup>* App t2 s2"
+ assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" "s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* s2"
+ shows "App t1 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2"
using a by (blast intro: Beta_App_cong_aux)
lemmas Beta_congs = Beta_Lam_cong Beta_App_cong
lemma One_implies_Beta_star:
- assumes a: "t \<longrightarrow>\<^isub>1 s"
- shows "t \<longrightarrow>\<^isub>\<beta>\<^sup>* s"
+ assumes a: "t \<longrightarrow>\<^sub>1 s"
+ shows "t \<longrightarrow>\<^sub>\<beta>\<^sup>* s"
using a by (induct) (auto intro!: Beta_congs)
lemma One_congs:
- assumes a: "t1 \<longrightarrow>\<^isub>1\<^sup>* t2"
- shows "Lam [x].t1 \<longrightarrow>\<^isub>1\<^sup>* Lam [x].t2"
- and "App t1 s \<longrightarrow>\<^isub>1\<^sup>* App t2 s"
- and "App s t1 \<longrightarrow>\<^isub>1\<^sup>* App s t2"
+ assumes a: "t1 \<longrightarrow>\<^sub>1\<^sup>* t2"
+ shows "Lam [x].t1 \<longrightarrow>\<^sub>1\<^sup>* Lam [x].t2"
+ and "App t1 s \<longrightarrow>\<^sub>1\<^sup>* App t2 s"
+ and "App s t1 \<longrightarrow>\<^sub>1\<^sup>* App s t2"
using a by (induct) (auto intro: One_refl)
lemma Beta_implies_One_star:
- assumes a: "t1 \<longrightarrow>\<^isub>\<beta> t2"
- shows "t1 \<longrightarrow>\<^isub>1\<^sup>* t2"
+ assumes a: "t1 \<longrightarrow>\<^sub>\<beta> t2"
+ shows "t1 \<longrightarrow>\<^sub>1\<^sup>* t2"
using a by (induct) (auto intro: One_refl One_congs better_o4_intro)
lemma Beta_star_equals_One_star:
- shows "t1 \<longrightarrow>\<^isub>1\<^sup>* t2 = t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2"
+ shows "t1 \<longrightarrow>\<^sub>1\<^sup>* t2 = t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
proof
- assume "t1 \<longrightarrow>\<^isub>1\<^sup>* t2"
- then show "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2" by (induct) (auto intro: One_implies_Beta_star)
+ assume "t1 \<longrightarrow>\<^sub>1\<^sup>* t2"
+ then show "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" by (induct) (auto intro: One_implies_Beta_star)
next
- assume "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2"
- then show "t1 \<longrightarrow>\<^isub>1\<^sup>* t2" by (induct) (auto intro: Beta_implies_One_star)
+ assume "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
+ then show "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" by (induct) (auto intro: Beta_implies_One_star)
qed
section {* The Church-Rosser Theorem *}
theorem CR_for_Beta_star:
- assumes a: "t \<longrightarrow>\<^isub>\<beta>\<^sup>* t1" "t\<longrightarrow>\<^isub>\<beta>\<^sup>* t2"
- shows "\<exists>t3. t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3"
+ assumes a: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t1" "t\<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
+ shows "\<exists>t3. t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3"
proof -
- from a have "t \<longrightarrow>\<^isub>1\<^sup>* t1" and "t\<longrightarrow>\<^isub>1\<^sup>* t2" by (simp_all add: Beta_star_equals_One_star)
- then have "\<exists>t3. t1 \<longrightarrow>\<^isub>1\<^sup>* t3 \<and> t2 \<longrightarrow>\<^isub>1\<^sup>* t3" by (simp add: CR_for_One_star)
- then show "\<exists>t3. t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^isub>\<beta>\<^sup>* t3" by (simp add: Beta_star_equals_One_star)
+ from a have "t \<longrightarrow>\<^sub>1\<^sup>* t1" and "t\<longrightarrow>\<^sub>1\<^sup>* t2" by (simp_all add: Beta_star_equals_One_star)
+ then have "\<exists>t3. t1 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by (simp add: CR_for_One_star)
+ then show "\<exists>t3. t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3" by (simp add: Beta_star_equals_One_star)
qed
--- a/src/HOL/Nominal/Examples/Class1.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Class1.thy Tue Aug 13 16:25:47 2013 +0200
@@ -3770,31 +3770,31 @@
done
inductive
- l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>l _" [100,100] 100)
+ l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>l _" [100,100] 100)
where
- LAxR: "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l M[a\<turnstile>c>b]"
-| LAxL: "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l M[x\<turnstile>n>y]"
+ LAxR: "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
+| LAxL: "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[x\<turnstile>n>y]"
| LNot: "\<lbrakk>y\<sharp>(M,N); x\<sharp>(N,y); a\<sharp>(M,N,b); b\<sharp>M; y\<noteq>x; b\<noteq>a\<rbrakk> \<Longrightarrow>
- Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.N (x).M"
+ Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).M"
| LAnd1: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
- Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <a1>.M1 (x).N"
+ Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <a1>.M1 (x).N"
| LAnd2: "\<lbrakk>b\<sharp>([a1].M1,[a2].M2,N,a1,a2); y\<sharp>([x].N,M1,M2,x); x\<sharp>(M1,M2); a1\<sharp>(M2,N); a2\<sharp>(M1,N); a1\<noteq>a2\<rbrakk> \<Longrightarrow>
- Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <a2>.M2 (x).N"
+ Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <a2>.M2 (x).N"
| LOr1: "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
- Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1"
+ Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
| LOr2: "\<lbrakk>b\<sharp>([a].M,N1,N2,a); y\<sharp>([x1].N1,[x2].N2,M,x1,x2); x1\<sharp>(M,N2); x2\<sharp>(M,N1); a\<sharp>(N1,N2); x1\<noteq>x2\<rbrakk> \<Longrightarrow>
- Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2"
+ Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
| LImp: "\<lbrakk>z\<sharp>(N,[y].P,[x].M,y,x); b\<sharp>([a].M,[c].N,P,c,a); x\<sharp>(N,[y].P,y);
c\<sharp>(P,[a].M,b,a); a\<sharp>([c].N,P); y\<sharp>(N,[x].M)\<rbrakk> \<Longrightarrow>
- Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
+ Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
equivariance l_redu
lemma l_redu_eqvt':
fixes pi1::"name prm"
and pi2::"coname prm"
- shows "(pi1\<bullet>M) \<longrightarrow>\<^isub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l M'"
- and "(pi2\<bullet>M) \<longrightarrow>\<^isub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^isub>l M'"
+ shows "(pi1\<bullet>M) \<longrightarrow>\<^sub>l (pi1\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l M'"
+ and "(pi2\<bullet>M) \<longrightarrow>\<^sub>l (pi2\<bullet>M') \<Longrightarrow> M \<longrightarrow>\<^sub>l M'"
apply -
apply(drule_tac pi="rev pi1" in l_redu.eqvt(1))
apply(perm_simp)
@@ -3810,8 +3810,8 @@
lemma fresh_l_redu:
fixes x::"name"
and a::"coname"
- shows "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
- and "M \<longrightarrow>\<^isub>l M' \<Longrightarrow> a\<sharp>M \<Longrightarrow> a\<sharp>M'"
+ shows "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
+ and "M \<longrightarrow>\<^sub>l M' \<Longrightarrow> a\<sharp>M \<Longrightarrow> a\<sharp>M'"
apply -
apply(induct rule: l_redu.induct)
apply(auto simp add: abs_fresh rename_fresh)
@@ -3828,35 +3828,35 @@
done
lemma better_LAxR_intro[intro]:
- shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>l M[a\<turnstile>c>b]"
+ shows "fic M a \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
proof -
assume fin: "fic M a"
obtain x'::"name" where fs1: "x'\<sharp>(M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(a,M,b)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.M (x).(Ax x b) = Cut <a'>.([(a',a)]\<bullet>M) (x').(Ax x' b)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 fin
+ also have "\<dots> \<longrightarrow>\<^sub>l ([(a',a)]\<bullet>M)[a'\<turnstile>c>b]" using fs1 fs2 fin
by (auto intro: l_redu.intros simp add: fresh_left calc_atm fic_rename)
also have "\<dots> = M[a\<turnstile>c>b]" using fs1 fs2 by (simp add: crename_rename)
finally show ?thesis by simp
qed
lemma better_LAxL_intro[intro]:
- shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>l M[x\<turnstile>n>y]"
+ shows "fin M x \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[x\<turnstile>n>y]"
proof -
assume fin: "fin M x"
obtain x'::"name" where fs1: "x'\<sharp>(y,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(a,M)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.(Ax y a) (x).M = Cut <a'>.(Ax y a') (x').([(x',x)]\<bullet>M)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 fin
+ also have "\<dots> \<longrightarrow>\<^sub>l ([(x',x)]\<bullet>M)[x'\<turnstile>n>y]" using fs1 fs2 fin
by (auto intro: l_redu.intros simp add: fresh_left calc_atm fin_rename)
also have "\<dots> = M[x\<turnstile>n>y]" using fs1 fs2 by (simp add: nrename_rename)
finally show ?thesis by simp
qed
lemma better_LNot_intro[intro]:
- shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>l Cut <b>.N (x).M"
+ shows "\<lbrakk>y\<sharp>N; a\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^sub>l Cut <b>.N (x).M"
proof -
assume fs: "y\<sharp>N" "a\<sharp>M"
obtain x'::"name" where f1: "x'\<sharp>(y,N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -3872,7 +3872,7 @@
also have "\<dots> = Cut <a'>.(NotR (x').([(x',x)]\<bullet>M) a') (y').(NotL <b'>.([(b',b)]\<bullet>N) y')"
using f1 f2 f3 f4
by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>M)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <b'>.([(b',b)]\<bullet>N) (x').([(x',x)]\<bullet>M)"
using f1 f2 f3 f4 fs
by (auto intro: l_redu.intros simp add: fresh_prod fresh_left calc_atm fresh_atm)
also have "\<dots> = Cut <b>.N (x).M"
@@ -3882,7 +3882,7 @@
lemma better_LAnd1_intro[intro]:
shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk>
- \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>l Cut <b1>.M1 (x).N"
+ \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL1 (x).N y) \<longrightarrow>\<^sub>l Cut <b1>.M1 (x).N"
proof -
assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[x].N"
obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -3903,7 +3903,7 @@
apply(rule_tac sym)
apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
done
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (x').([(x',x)]\<bullet>N)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <b1'>.([(b1',b1)]\<bullet>M1) (x').([(x',x)]\<bullet>N)"
using f1 f2 f3 f4 f5 fs
apply -
apply(rule l_redu.intros)
@@ -3916,7 +3916,7 @@
lemma better_LAnd2_intro[intro]:
shows "\<lbrakk>a\<sharp>([b1].M1,[b2].M2); y\<sharp>[x].N\<rbrakk>
- \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>l Cut <b2>.M2 (x).N"
+ \<Longrightarrow> Cut <a>.(AndR <b1>.M1 <b2>.M2 a) (y).(AndL2 (x).N y) \<longrightarrow>\<^sub>l Cut <b2>.M2 (x).N"
proof -
assume fs: "a\<sharp>([b1].M1,[b2].M2)" "y\<sharp>[x].N"
obtain x'::"name" where f1: "x'\<sharp>(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -3937,7 +3937,7 @@
apply(rule_tac sym)
apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
done
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (x').([(x',x)]\<bullet>N)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <b2'>.([(b2',b2)]\<bullet>M2) (x').([(x',x)]\<bullet>N)"
using f1 f2 f3 f4 f5 fs
apply -
apply(rule l_redu.intros)
@@ -3950,7 +3950,7 @@
lemma better_LOr1_intro[intro]:
shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk>
- \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x1).N1"
+ \<Longrightarrow> Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x1).N1"
proof -
assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].M"
obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -3971,7 +3971,7 @@
apply(rule_tac sym)
apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
done
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x1').([(x1',x1)]\<bullet>N1)"
using f1 f2 f3 f4 f5 fs
apply -
apply(rule l_redu.intros)
@@ -3984,7 +3984,7 @@
lemma better_LOr2_intro[intro]:
shows "\<lbrakk>y\<sharp>([x1].N1,[x2].N2); b\<sharp>[a].M\<rbrakk>
- \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>l Cut <a>.M (x2).N2"
+ \<Longrightarrow> Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^sub>l Cut <a>.M (x2).N2"
proof -
assume fs: "y\<sharp>([x1].N1,[x2].N2)" "b\<sharp>[a].M"
obtain y'::"name" where f1: "y'\<sharp>(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -4005,7 +4005,7 @@
apply(rule_tac sym)
apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh)
done
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.([(a',a)]\<bullet>M) (x2').([(x2',x2)]\<bullet>N2)"
using f1 f2 f3 f4 f5 fs
apply -
apply(rule l_redu.intros)
@@ -4018,7 +4018,7 @@
lemma better_LImp_intro[intro]:
shows "\<lbrakk>z\<sharp>(N,[y].P); b\<sharp>[a].M; a\<sharp>N\<rbrakk>
- \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
+ \<Longrightarrow> Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^sub>l Cut <a>.(Cut <c>.N (x).M) (y).P"
proof -
assume fs: "z\<sharp>(N,[y].P)" "b\<sharp>[a].M" "a\<sharp>N"
obtain y'::"name" where f1: "y'\<sharp>(y,M,N,P,z,x)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -4049,7 +4049,7 @@
apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh)
done
- also have "\<dots> \<longrightarrow>\<^isub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<bullet>P)"
+ also have "\<dots> \<longrightarrow>\<^sub>l Cut <a'>.(Cut <c'>.([(c',c)]\<bullet>N) (x').([(a',a)]\<bullet>[(x',x)]\<bullet>M)) (y').([(y',y)]\<bullet>P)"
using f1 f2 f3 f4 f5 f6 fs
apply -
apply(rule l_redu.intros)
@@ -4118,7 +4118,7 @@
done
lemma Cut_l_redu_elim:
- assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>l R"
+ assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>l R"
shows "(\<exists>b. R = M[a\<turnstile>c>b]) \<or> (\<exists>y. R = N[x\<turnstile>n>y]) \<or>
(\<exists>y M' b N'. M = NotR (y).M' a \<and> N = NotL <b>.N' x \<and> R = Cut <b>.N' (y).M' \<and> fic M a \<and> fin N x) \<or>
(\<exists>b M1 c M2 y N'. M = AndR <b>.M1 <c>.M2 a \<and> N = AndL1 (y).N' x \<and> R = Cut <b>.M1 (y).N'
@@ -4471,10 +4471,10 @@
done
inductive
- c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>c _" [100,100] 100)
+ c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>c _" [100,100] 100)
where
- left[intro]: "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(x).N}"
-| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<a>.M}"
+ left[intro]: "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
+| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<a>.M}"
equivariance c_redu
@@ -4482,14 +4482,14 @@
by (simp_all add: abs_fresh subst_fresh)
lemma better_left[intro]:
- shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M{a:=(x).N}"
+ shows "\<not>fic M a \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
proof -
assume not_fic: "\<not>fic M a"
obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_fic
+ also have "\<dots> \<longrightarrow>\<^sub>c ([(a',a)]\<bullet>M){a':=(x').([(x',x)]\<bullet>N)}" using fs1 fs2 not_fic
apply -
apply(rule left)
apply(clarify)
@@ -4503,14 +4503,14 @@
qed
lemma better_right[intro]:
- shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N{x:=<a>.M}"
+ shows "\<not>fin N x \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<a>.M}"
proof -
assume not_fin: "\<not>fin N x"
obtain x'::"name" where fs1: "x'\<sharp>(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_fin
+ also have "\<dots> \<longrightarrow>\<^sub>c ([(x',x)]\<bullet>N){x':=<a'>.([(a',a)]\<bullet>M)}" using fs1 fs2 not_fin
apply -
apply(rule right)
apply(clarify)
@@ -4526,8 +4526,8 @@
lemma fresh_c_redu:
fixes x::"name"
and c::"coname"
- shows "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
- and "M \<longrightarrow>\<^isub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
+ shows "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
+ and "M \<longrightarrow>\<^sub>c M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
apply -
apply(induct rule: c_redu.induct)
apply(auto simp add: abs_fresh rename_fresh subst_fresh)
@@ -4536,31 +4536,31 @@
done
inductive
- a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a _" [100,100] 100)
+ a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a _" [100,100] 100)
where
- al_redu[intro]: "M\<longrightarrow>\<^isub>l M' \<Longrightarrow> M \<longrightarrow>\<^isub>a M'"
-| ac_redu[intro]: "M\<longrightarrow>\<^isub>c M' \<Longrightarrow> M \<longrightarrow>\<^isub>a M'"
-| a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N"
-| a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'"
-| a_NotL[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a NotL <a>.M' x"
-| a_NotR[intro]: "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a NotR (x).M' a"
-| a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c"
-| a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c"
-| a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).M' y"
-| a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).M' y"
-| a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z"
-| a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z"
-| a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.M' b"
-| a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.M' b"
-| a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (x).N y"
-| a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^isub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (x).N' y"
-| a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^isub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.M' b"
+ al_redu[intro]: "M\<longrightarrow>\<^sub>l M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
+| ac_redu[intro]: "M\<longrightarrow>\<^sub>c M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
+| a_Cut_l: "\<lbrakk>a\<sharp>N; x\<sharp>M; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).N"
+| a_Cut_r: "\<lbrakk>a\<sharp>N; x\<sharp>M; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).N'"
+| a_NotL[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a NotL <a>.M' x"
+| a_NotR[intro]: "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a NotR (x).M' a"
+| a_AndR_l: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N c"
+| a_AndR_r: "\<lbrakk>a\<sharp>(N,c); b\<sharp>(M,c); b\<noteq>a; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' c"
+| a_AndL1: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' y"
+| a_AndL2: "\<lbrakk>x\<sharp>y; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' y"
+| a_OrL_l: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).N z"
+| a_OrL_r: "\<lbrakk>x\<sharp>(N,z); y\<sharp>(M,z); y\<noteq>x; N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).N' z"
+| a_OrR1: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' b"
+| a_OrR2: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' b"
+| a_ImpL_l: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).N y"
+| a_ImpL_r: "\<lbrakk>a\<sharp>N; x\<sharp>(M,y); N\<longrightarrow>\<^sub>a N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).N' y"
+| a_ImpR: "\<lbrakk>a\<sharp>b; M\<longrightarrow>\<^sub>a M'\<rbrakk> \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' b"
lemma fresh_a_redu:
fixes x::"name"
and c::"coname"
- shows "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
- and "M \<longrightarrow>\<^isub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
+ shows "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> x\<sharp>M \<Longrightarrow> x\<sharp>M'"
+ and "M \<longrightarrow>\<^sub>a M' \<Longrightarrow> c\<sharp>M \<Longrightarrow> c\<sharp>M'"
apply -
apply(induct rule: a_redu.induct)
apply(simp add: fresh_l_redu)
@@ -4578,14 +4578,14 @@
by (simp_all add: abs_fresh fresh_atm fresh_prod abs_supp fin_supp fresh_a_redu)
lemma better_CutL_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M' (x).N"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M' (x).N"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a Cut <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N)" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt)
also have "\<dots> = Cut <a>.M' (x).N"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4593,14 +4593,14 @@
qed
lemma better_CutR_intro[intro]:
- shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a Cut <a>.M (x).N'"
+ shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a Cut <a>.M (x).N'"
proof -
- assume red: "N\<longrightarrow>\<^isub>a N'"
+ assume red: "N\<longrightarrow>\<^sub>a N'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "Cut <a>.M (x).N = Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N)"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a Cut <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N')" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt)
also have "\<dots> = Cut <a>.M (x).N'"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4608,14 +4608,14 @@
qed
lemma better_AndRL_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M' <b>.N c"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M' <b>.N c"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast)
have "AndR <a>.M <b>.N c = AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a AndR <a'>.([(a',a)]\<bullet>M') <b'>.([(b',b)]\<bullet>N) c" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = AndR <a>.M' <b>.N c"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4623,14 +4623,14 @@
qed
lemma better_AndRR_intro[intro]:
- shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a AndR <a>.M <b>.N' c"
+ shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a AndR <a>.M <b>.N' c"
proof -
- assume red: "N\<longrightarrow>\<^isub>a N'"
+ assume red: "N\<longrightarrow>\<^sub>a N'"
obtain b'::"coname" where fs1: "b'\<sharp>(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast)
have "AndR <a>.M <b>.N c = AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N) c"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a AndR <a'>.([(a',a)]\<bullet>M) <b'>.([(b',b)]\<bullet>N') c" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = AndR <a>.M <b>.N' c"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4638,13 +4638,13 @@
qed
lemma better_AndL1_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a AndL1 (x).M' y"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a AndL1 (x).M' y"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast)
have "AndL1 (x).M y = AndL1 (x').([(x',x)]\<bullet>M) y"
using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red
+ also have "\<dots> \<longrightarrow>\<^sub>a AndL1 (x').([(x',x)]\<bullet>M') y" using fs1 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = AndL1 (x).M' y"
using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4652,13 +4652,13 @@
qed
lemma better_AndL2_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a AndL2 (x).M' y"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a AndL2 (x).M' y"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain x'::"name" where fs1: "x'\<sharp>(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast)
have "AndL2 (x).M y = AndL2 (x').([(x',x)]\<bullet>M) y"
using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red
+ also have "\<dots> \<longrightarrow>\<^sub>a AndL2 (x').([(x',x)]\<bullet>M') y" using fs1 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = AndL2 (x).M' y"
using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4666,14 +4666,14 @@
qed
lemma better_OrLL_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M' (y).N z"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M' (y).N z"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast)
have "OrL (x).M (y).N z = OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M') (y').([(y',y)]\<bullet>N) z" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = OrL (x).M' (y).N z"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4681,14 +4681,14 @@
qed
lemma better_OrLR_intro[intro]:
- shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a OrL (x).M (y).N' z"
+ shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a OrL (x).M (y).N' z"
proof -
- assume red: "N\<longrightarrow>\<^isub>a N'"
+ assume red: "N\<longrightarrow>\<^sub>a N'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain y'::"name" where fs2: "y'\<sharp>(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast)
have "OrL (x).M (y).N z = OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N) z"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a OrL (x').([(x',x)]\<bullet>M) (y').([(y',y)]\<bullet>N') z" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = OrL (x).M (y).N' z"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4696,13 +4696,13 @@
qed
lemma better_OrR1_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a OrR1 <a>.M' b"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a OrR1 <a>.M' b"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "OrR1 <a>.M b = OrR1 <a'>.([(a',a)]\<bullet>M) b"
using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
+ also have "\<dots> \<longrightarrow>\<^sub>a OrR1 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = OrR1 <a>.M' b"
using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4710,13 +4710,13 @@
qed
lemma better_OrR2_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a OrR2 <a>.M' b"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a OrR2 <a>.M' b"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain a'::"coname" where fs1: "a'\<sharp>(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "OrR2 <a>.M b = OrR2 <a'>.([(a',a)]\<bullet>M) b"
using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
+ also have "\<dots> \<longrightarrow>\<^sub>a OrR2 <a'>.([(a',a)]\<bullet>M') b" using fs1 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = OrR2 <a>.M' b"
using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4724,14 +4724,14 @@
qed
lemma better_ImpLL_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M' (x).N y"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M' (x).N y"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "ImpL <a>.M (x).N y = ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a ImpL <a'>.([(a',a)]\<bullet>M') (x').([(x',x)]\<bullet>N) y" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = ImpL <a>.M' (x).N y"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4739,14 +4739,14 @@
qed
lemma better_ImpLR_intro[intro]:
- shows "N\<longrightarrow>\<^isub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^isub>a ImpL <a>.M (x).N' y"
+ shows "N\<longrightarrow>\<^sub>a N' \<Longrightarrow> ImpL <a>.M (x).N y \<longrightarrow>\<^sub>a ImpL <a>.M (x).N' y"
proof -
- assume red: "N\<longrightarrow>\<^isub>a N'"
+ assume red: "N\<longrightarrow>\<^sub>a N'"
obtain x'::"name" where fs1: "x'\<sharp>(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast)
obtain a'::"coname" where fs2: "a'\<sharp>(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast)
have "ImpL <a>.M (x).N y = ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N) y"
using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a ImpL <a'>.([(a',a)]\<bullet>M) (x').([(x',x)]\<bullet>N') y" using fs1 fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = ImpL <a>.M (x).N' y"
using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4754,13 +4754,13 @@
qed
lemma better_ImpR_intro[intro]:
- shows "M\<longrightarrow>\<^isub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a ImpR (x).<a>.M' b"
+ shows "M\<longrightarrow>\<^sub>a M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a ImpR (x).<a>.M' b"
proof -
- assume red: "M\<longrightarrow>\<^isub>a M'"
+ assume red: "M\<longrightarrow>\<^sub>a M'"
obtain a'::"coname" where fs2: "a'\<sharp>(M,a,b)" by (rule exists_fresh(2), rule fin_supp, blast)
have "ImpR (x).<a>.M b = ImpR (x).<a'>.([(a',a)]\<bullet>M) b"
using fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red
+ also have "\<dots> \<longrightarrow>\<^sub>a ImpR (x).<a'>.([(a',a)]\<bullet>M') b" using fs2 red
by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod)
also have "\<dots> = ImpR (x).<a>.M' b"
using fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu)
@@ -4770,15 +4770,15 @@
text {* axioms do not reduce *}
lemma ax_do_not_l_reduce:
- shows "Ax x a \<longrightarrow>\<^isub>l M \<Longrightarrow> False"
+ shows "Ax x a \<longrightarrow>\<^sub>l M \<Longrightarrow> False"
by (erule_tac l_redu.cases) (simp_all add: trm.inject)
lemma ax_do_not_c_reduce:
- shows "Ax x a \<longrightarrow>\<^isub>c M \<Longrightarrow> False"
+ shows "Ax x a \<longrightarrow>\<^sub>c M \<Longrightarrow> False"
by (erule_tac c_redu.cases) (simp_all add: trm.inject)
lemma ax_do_not_a_reduce:
- shows "Ax x a \<longrightarrow>\<^isub>a M \<Longrightarrow> False"
+ shows "Ax x a \<longrightarrow>\<^sub>a M \<Longrightarrow> False"
apply(erule_tac a_redu.cases)
apply(auto simp add: trm.inject)
apply(drule ax_do_not_l_reduce)
@@ -4788,8 +4788,8 @@
done
lemma a_redu_NotL_elim:
- assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "NotL <a>.M x \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = NotL <a>.M' x \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -4809,8 +4809,8 @@
done
lemma a_redu_NotR_elim:
- assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "NotR (x).M a \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = NotR (x).M' a \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -4830,8 +4830,8 @@
done
lemma a_redu_AndR_elim:
- assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a R"
- shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^isub>aN')"
+ assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a R"
+ shows "(\<exists>M'. R = AndR <a>.M' <b>.N c \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = AndR <a>.M <b>.N' c \<and> N\<longrightarrow>\<^sub>aN')"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -4919,8 +4919,8 @@
done
lemma a_redu_AndL1_elim:
- assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "AndL1 (x).M y \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = AndL1 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -4940,8 +4940,8 @@
done
lemma a_redu_AndL2_elim:
- assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "AndL2 (x).M y \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = AndL2 (x).M' y \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -4961,8 +4961,8 @@
done
lemma a_redu_OrL_elim:
- assumes a: "OrL (x).M (y).N z\<longrightarrow>\<^isub>a R"
- shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')"
+ assumes a: "OrL (x).M (y).N z\<longrightarrow>\<^sub>a R"
+ shows "(\<exists>M'. R = OrL (x).M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = OrL (x).M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -5050,8 +5050,8 @@
done
lemma a_redu_OrR1_elim:
- assumes a: "OrR1 <a>.M b \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "OrR1 <a>.M b \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = OrR1 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -5071,8 +5071,8 @@
done
lemma a_redu_OrR2_elim:
- assumes a: "OrR2 <a>.M b \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "OrR2 <a>.M b \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = OrR2 <a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -5092,8 +5092,8 @@
done
lemma a_redu_ImpL_elim:
- assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^isub>a R"
- shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^isub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^isub>aN')"
+ assumes a: "ImpL <a>.M (y).N z\<longrightarrow>\<^sub>a R"
+ shows "(\<exists>M'. R = ImpL <a>.M' (y).N z \<and> M\<longrightarrow>\<^sub>aM') \<or> (\<exists>N'. R = ImpL <a>.M (y).N' z \<and> N\<longrightarrow>\<^sub>aN')"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -5181,8 +5181,8 @@
done
lemma a_redu_ImpR_elim:
- assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^isub>a R"
- shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^isub>aM'"
+ assumes a: "ImpR (x).<a>.M b \<longrightarrow>\<^sub>a R"
+ shows "\<exists>M'. R = ImpR (x).<a>.M' b \<and> M\<longrightarrow>\<^sub>aM'"
using a
apply(erule_tac a_redu.cases, simp_all add: trm.inject)
apply(erule_tac l_redu.cases, simp_all add: trm.inject)
@@ -5238,36 +5238,36 @@
text {* Transitive Closure*}
abbreviation
- a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>a* _" [100,100] 100)
+ a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a* _" [100,100] 100)
where
- "M \<longrightarrow>\<^isub>a* M' \<equiv> (a_redu)^** M M'"
+ "M \<longrightarrow>\<^sub>a* M' \<equiv> (a_redu)^** M M'"
lemma a_starI:
- assumes a: "M \<longrightarrow>\<^isub>a M'"
- shows "M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "M \<longrightarrow>\<^sub>a M'"
+ shows "M \<longrightarrow>\<^sub>a* M'"
using a by blast
lemma a_starE:
- assumes a: "M \<longrightarrow>\<^isub>a* M'"
- shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^isub>a N \<and> N \<longrightarrow>\<^isub>a* M')"
+ assumes a: "M \<longrightarrow>\<^sub>a* M'"
+ shows "M = M' \<or> (\<exists>N. M \<longrightarrow>\<^sub>a N \<and> N \<longrightarrow>\<^sub>a* M')"
using a
by (induct) (auto)
lemma a_star_refl:
- shows "M \<longrightarrow>\<^isub>a* M"
+ shows "M \<longrightarrow>\<^sub>a* M"
by blast
lemma a_star_trans[trans]:
- assumes a1: "M1\<longrightarrow>\<^isub>a* M2"
- and a2: "M2\<longrightarrow>\<^isub>a* M3"
- shows "M1 \<longrightarrow>\<^isub>a* M3"
+ assumes a1: "M1\<longrightarrow>\<^sub>a* M2"
+ and a2: "M2\<longrightarrow>\<^sub>a* M3"
+ shows "M1 \<longrightarrow>\<^sub>a* M3"
using a2 a1
by (induct) (auto)
-text {* congruence rules for \<longrightarrow>\<^isub>a* *}
+text {* congruence rules for \<longrightarrow>\<^sub>a* *}
lemma ax_do_not_a_star_reduce:
- shows "Ax x a \<longrightarrow>\<^isub>a* M \<Longrightarrow> M = Ax x a"
+ shows "Ax x a \<longrightarrow>\<^sub>a* M \<Longrightarrow> M = Ax x a"
apply(induct set: rtranclp)
apply(auto)
apply(drule ax_do_not_a_reduce)
@@ -5276,87 +5276,87 @@
lemma a_star_CutL:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).N"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_CutR:
- "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M (x).N'"
+ "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M (x).N'"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_Cut:
- "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>a* Cut <a>.M' (x).N'"
+ "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>a* Cut <a>.M' (x).N'"
by (blast intro!: a_star_CutL a_star_CutR intro: rtranclp_trans)
lemma a_star_NotR:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^isub>a* NotR (x).M' a"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotR (x).M a \<longrightarrow>\<^sub>a* NotR (x).M' a"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_NotL:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^isub>a* NotL <a>.M' x"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> NotL <a>.M x \<longrightarrow>\<^sub>a* NotL <a>.M' x"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_AndRL:
- "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N c"
+ "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N c"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_AndRR:
- "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M <b>.N' c"
+ "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M <b>.N' c"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_AndR:
- "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^isub>a* AndR <a>.M' <b>.N' c"
+ "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> AndR <a>.M <b>.N c \<longrightarrow>\<^sub>a* AndR <a>.M' <b>.N' c"
by (blast intro!: a_star_AndRL a_star_AndRR intro: rtranclp_trans)
lemma a_star_AndL1:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^isub>a* AndL1 (x).M' y"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL1 (x).M y \<longrightarrow>\<^sub>a* AndL1 (x).M' y"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_AndL2:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^isub>a* AndL2 (x).M' y"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> AndL2 (x).M y \<longrightarrow>\<^sub>a* AndL2 (x).M' y"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_OrLL:
- "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N z"
+ "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).N z"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_OrLR:
- "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M (y).N' z"
+ "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M (y).N' z"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_OrL:
- "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^isub>a* OrL (x).M' (y).N' z"
+ "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> OrL (x).M (y).N z \<longrightarrow>\<^sub>a* OrL (x).M' (y).N' z"
by (blast intro!: a_star_OrLL a_star_OrLR intro: rtranclp_trans)
lemma a_star_OrR1:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^isub>a* OrR1 <a>.M' b"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR1 <a>.M b \<longrightarrow>\<^sub>a* OrR1 <a>.M' b"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_OrR2:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^isub>a* OrR2 <a>.M' b"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> OrR2 <a>.M b \<longrightarrow>\<^sub>a* OrR2 <a>.M' b"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_ImpLL:
- "M \<longrightarrow>\<^isub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N z"
+ "M \<longrightarrow>\<^sub>a* M'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).N z"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_ImpLR:
- "N \<longrightarrow>\<^isub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M (y).N' z"
+ "N \<longrightarrow>\<^sub>a* N'\<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M (y).N' z"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemma a_star_ImpL:
- "\<lbrakk>M \<longrightarrow>\<^isub>a* M'; N \<longrightarrow>\<^isub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* ImpL <a>.M' (y).N' z"
+ "\<lbrakk>M \<longrightarrow>\<^sub>a* M'; N \<longrightarrow>\<^sub>a* N'\<rbrakk> \<Longrightarrow> ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* ImpL <a>.M' (y).N' z"
by (blast intro!: a_star_ImpLL a_star_ImpLR intro: rtranclp_trans)
lemma a_star_ImpR:
- "M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^isub>a* ImpR (x).<a>.M' b"
+ "M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> ImpR (x).<a>.M b \<longrightarrow>\<^sub>a* ImpR (x).<a>.M' b"
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+
lemmas a_star_congs = a_star_Cut a_star_NotR a_star_NotL a_star_AndR a_star_AndL1 a_star_AndL2
a_star_OrL a_star_OrR1 a_star_OrR2 a_star_ImpL a_star_ImpR
lemma a_star_redu_NotL_elim:
- assumes a: "NotL <a>.M x \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = NotL <a>.M' x \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "NotL <a>.M x \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = NotL <a>.M' x \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5365,8 +5365,8 @@
done
lemma a_star_redu_NotR_elim:
- assumes a: "NotR (x).M a \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = NotR (x).M' a \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "NotR (x).M a \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = NotR (x).M' a \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5375,8 +5375,8 @@
done
lemma a_star_redu_AndR_elim:
- assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^isub>a* R"
- shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
+ assumes a: "AndR <a>.M <b>.N c\<longrightarrow>\<^sub>a* R"
+ shows "(\<exists>M' N'. R = AndR <a>.M' <b>.N' c \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5385,8 +5385,8 @@
done
lemma a_star_redu_AndL1_elim:
- assumes a: "AndL1 (x).M y \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = AndL1 (x).M' y \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "AndL1 (x).M y \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = AndL1 (x).M' y \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5395,8 +5395,8 @@
done
lemma a_star_redu_AndL2_elim:
- assumes a: "AndL2 (x).M y \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = AndL2 (x).M' y \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "AndL2 (x).M y \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = AndL2 (x).M' y \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5405,8 +5405,8 @@
done
lemma a_star_redu_OrL_elim:
- assumes a: "OrL (x).M (y).N z \<longrightarrow>\<^isub>a* R"
- shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
+ assumes a: "OrL (x).M (y).N z \<longrightarrow>\<^sub>a* R"
+ shows "(\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5415,8 +5415,8 @@
done
lemma a_star_redu_OrR1_elim:
- assumes a: "OrR1 <a>.M y \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = OrR1 <a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "OrR1 <a>.M y \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = OrR1 <a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5425,8 +5425,8 @@
done
lemma a_star_redu_OrR2_elim:
- assumes a: "OrR2 <a>.M y \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = OrR2 <a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "OrR2 <a>.M y \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = OrR2 <a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5435,8 +5435,8 @@
done
lemma a_star_redu_ImpR_elim:
- assumes a: "ImpR (x).<a>.M y \<longrightarrow>\<^isub>a* R"
- shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> M \<longrightarrow>\<^isub>a* M'"
+ assumes a: "ImpR (x).<a>.M y \<longrightarrow>\<^sub>a* R"
+ shows "\<exists>M'. R = ImpR (x).<a>.M' y \<and> M \<longrightarrow>\<^sub>a* M'"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5445,8 +5445,8 @@
done
lemma a_star_redu_ImpL_elim:
- assumes a: "ImpL <a>.M (y).N z \<longrightarrow>\<^isub>a* R"
- shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^isub>a* M' \<and> N \<longrightarrow>\<^isub>a* N')"
+ assumes a: "ImpL <a>.M (y).N z \<longrightarrow>\<^sub>a* R"
+ shows "(\<exists>M' N'. R = ImpL <a>.M' (y).N' z \<and> M \<longrightarrow>\<^sub>a* M' \<and> N \<longrightarrow>\<^sub>a* N')"
using a
apply(induct set: rtranclp)
apply(auto)
@@ -5723,7 +5723,7 @@
lemma fin_l_reduce:
assumes a: "fin M x"
- and b: "M \<longrightarrow>\<^isub>l M'"
+ and b: "M \<longrightarrow>\<^sub>l M'"
shows "fin M' x"
using b a
apply(induct)
@@ -5737,7 +5737,7 @@
lemma fin_c_reduce:
assumes a: "fin M x"
- and b: "M \<longrightarrow>\<^isub>c M'"
+ and b: "M \<longrightarrow>\<^sub>c M'"
shows "fin M' x"
using b a
apply(induct)
@@ -5746,7 +5746,7 @@
lemma fin_a_reduce:
assumes a: "fin M x"
- and b: "M \<longrightarrow>\<^isub>a M'"
+ and b: "M \<longrightarrow>\<^sub>a M'"
shows "fin M' x"
using a b
apply(induct)
@@ -5784,7 +5784,7 @@
lemma fin_a_star_reduce:
assumes a: "fin M x"
- and b: "M \<longrightarrow>\<^isub>a* M'"
+ and b: "M \<longrightarrow>\<^sub>a* M'"
shows "fin M' x"
using b a
apply(induct set: rtranclp)
@@ -5793,7 +5793,7 @@
lemma fic_l_reduce:
assumes a: "fic M x"
- and b: "M \<longrightarrow>\<^isub>l M'"
+ and b: "M \<longrightarrow>\<^sub>l M'"
shows "fic M' x"
using b a
apply(induct)
@@ -5807,7 +5807,7 @@
lemma fic_c_reduce:
assumes a: "fic M x"
- and b: "M \<longrightarrow>\<^isub>c M'"
+ and b: "M \<longrightarrow>\<^sub>c M'"
shows "fic M' x"
using b a
apply(induct)
@@ -5816,7 +5816,7 @@
lemma fic_a_reduce:
assumes a: "fic M x"
- and b: "M \<longrightarrow>\<^isub>a M'"
+ and b: "M \<longrightarrow>\<^sub>a M'"
shows "fic M' x"
using a b
apply(induct)
@@ -5850,7 +5850,7 @@
lemma fic_a_star_reduce:
assumes a: "fic M x"
- and b: "M \<longrightarrow>\<^isub>a* M'"
+ and b: "M \<longrightarrow>\<^sub>a* M'"
shows "fic M' x"
using b a
apply(induct set: rtranclp)
@@ -5860,55 +5860,55 @@
text {* substitution properties *}
lemma subst_with_ax1:
- shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]"
+ shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]"
proof(nominal_induct M avoiding: x a y rule: trm.strong_induct)
case (Ax z b x a y)
- show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]"
+ show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]"
proof (cases "z=x")
case True
assume eq: "z=x"
have "(Ax z b){x:=<a>.Ax y a} = Cut <a>.Ax y a (x).Ax x b" using eq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* (Ax x b)[x\<turnstile>n>y]" by blast
- finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" using eq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Ax x b)[x\<turnstile>n>y]" by blast
+ finally show "Ax z b{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using eq by simp
next
case False
assume neq: "z\<noteq>x"
- then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Ax z b)[x\<turnstile>n>y]" using neq by simp
+ then show "(Ax z b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Ax z b)[x\<turnstile>n>y]" using neq by simp
qed
next
case (Cut b M z N x a y)
have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>N" "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>M" by fact+
- have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
- show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]"
+ have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
+ show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]"
proof (cases "M = Ax x b")
case True
assume eq: "M = Ax x b"
have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <a>.Ax y a (z).(N{x:=<a>.Ax y a})" using fs eq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs by blast
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.Ax y a (z).(N[x\<turnstile>n>y])" using ih2 a_star_congs by blast
also have "\<dots> = Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using eq
by (simp add: trm.inject alpha calc_atm fresh_atm)
- finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
+ finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
next
case False
assume neq: "M \<noteq> Ax x b"
have "(Cut <b>.M (z).N){x:=<a>.Ax y a} = Cut <b>.(M{x:=<a>.Ax y a}) (z).(N{x:=<a>.Ax y a})"
using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs by blast
- finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <b>.(M[x\<turnstile>n>y]) (z).(N[x\<turnstile>n>y])" using ih1 ih2 a_star_congs by blast
+ finally show "(Cut <b>.M (z).N){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (Cut <b>.M (z).N)[x\<turnstile>n>y]" using fs by simp
qed
next
case (NotR z M b x a y)
have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "z\<sharp>b" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
have "(NotR (z).M b){x:=<a>.Ax y a} = NotR (z).(M{x:=<a>.Ax y a}) b" using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih by (auto intro: a_star_congs)
- finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotR (z).M b)[x\<turnstile>n>y]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[x\<turnstile>n>y]) b" using ih by (auto intro: a_star_congs)
+ finally show "(NotR (z).M b){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotR (z).M b)[x\<turnstile>n>y]" using fs by simp
next
case (NotL b M z x a y)
have fs: "b\<sharp>x" "b\<sharp>a" "b\<sharp>y" "b\<sharp>z" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]"
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]"
proof(cases "z=x")
case True
assume eq: "z=x"
@@ -5918,7 +5918,7 @@
using eq fs by simp
also have "\<dots> = Cut <a>.Ax y a (x').NotL <b>.(M{x:=<a>.Ax y a}) x'"
using new by (simp add: fresh_fun_simp_NotL fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[x'\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (NotL <b>.(M{x:=<a>.Ax y a}) x')[x'\<turnstile>n>y]"
using new
apply(rule_tac a_starI)
apply(rule al_redu)
@@ -5926,31 +5926,31 @@
apply(auto)
done
also have "\<dots> = NotL <b>.(M{x:=<a>.Ax y a}) y" using new by (simp add: nrename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
also have "\<dots> = (NotL <b>.M z)[x\<turnstile>n>y]" using eq by simp
- finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" by simp
+ finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" by simp
next
case False
assume neq: "z\<noteq>x"
have "(NotL <b>.M z){x:=<a>.Ax y a} = NotL <b>.(M{x:=<a>.Ax y a}) z" using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih by (auto intro: a_star_congs)
- finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotL <b>.(M[x\<turnstile>n>y]) z" using ih by (auto intro: a_star_congs)
+ finally show "(NotL <b>.M z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (NotL <b>.M z)[x\<turnstile>n>y]" using neq by simp
qed
next
case (AndR c M d N e x a y)
have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "d\<sharp>x" "d\<sharp>a" "d\<sharp>y" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
- have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
+ have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
have "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} = AndR <c>.(M{x:=<a>.Ax y a}) <d>.(N{x:=<a>.Ax y a}) e"
using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[x\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[x\<turnstile>n>y]) <d>.(N[x\<turnstile>n>y]) e" using ih1 ih2 by (auto intro: a_star_congs)
+ finally show "(AndR <c>.M <d>.N e){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[x\<turnstile>n>y]"
using fs by simp
next
case (AndL1 u M v x a y)
have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]"
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]"
proof(cases "v=x")
case True
assume eq: "v=x"
@@ -5960,7 +5960,7 @@
using eq fs by simp
also have "\<dots> = Cut <a>.Ax y a (v').AndL1 (u).(M{x:=<a>.Ax y a}) v'"
using new by (simp add: fresh_fun_simp_AndL1 fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (AndL1 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -5970,21 +5970,21 @@
done
also have "\<dots> = AndL1 (u).(M{x:=<a>.Ax y a}) y" using fs new
by (auto simp add: fresh_prod fresh_atm nrename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
also have "\<dots> = (AndL1 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
- finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" by simp
+ finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" by simp
next
case False
assume neq: "v\<noteq>x"
have "(AndL1 (u).M v){x:=<a>.Ax y a} = AndL1 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
- finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
+ finally show "(AndL1 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
qed
next
case (AndL2 u M v x a y)
have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "u\<sharp>v" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]"
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]"
proof(cases "v=x")
case True
assume eq: "v=x"
@@ -5994,7 +5994,7 @@
using eq fs by simp
also have "\<dots> = Cut <a>.Ax y a (v').AndL2 (u).(M{x:=<a>.Ax y a}) v'"
using new by (simp add: fresh_fun_simp_AndL2 fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (AndL2 (u).(M{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6004,36 +6004,36 @@
done
also have "\<dots> = AndL2 (u).(M{x:=<a>.Ax y a}) y" using fs new
by (auto simp add: fresh_prod fresh_atm nrename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) y" using ih by (auto intro: a_star_congs)
also have "\<dots> = (AndL2 (u).M v)[x\<turnstile>n>y]" using eq fs by simp
- finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" by simp
+ finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" by simp
next
case False
assume neq: "v\<noteq>x"
have "(AndL2 (u).M v){x:=<a>.Ax y a} = AndL2 (u).(M{x:=<a>.Ax y a}) v" using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
- finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[x\<turnstile>n>y]) v" using ih by (auto intro: a_star_congs)
+ finally show "(AndL2 (u).M v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[x\<turnstile>n>y]" using fs neq by simp
qed
next
case (OrR1 c M d x a y)
have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
have "(OrR1 <c>.M d){x:=<a>.Ax y a} = OrR1 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
- finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
+ finally show "(OrR1 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[x\<turnstile>n>y]" using fs by simp
next
case (OrR2 c M d x a y)
have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "c\<sharp>d" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
have "(OrR2 <c>.M d){x:=<a>.Ax y a} = OrR2 <c>.(M{x:=<a>.Ax y a}) d" using fs by (simp add: fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
- finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
+ finally show "(OrR2 <c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[x\<turnstile>n>y]" using fs by simp
next
case (OrL u M v N z x a y)
have fs: "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "v\<sharp>x" "v\<sharp>a" "v\<sharp>y" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
- have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
- show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]"
+ have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
+ show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]"
proof(cases "z=x")
case True
assume eq: "z=x"
@@ -6044,7 +6044,7 @@
using eq fs by simp
also have "\<dots> = Cut <a>.Ax y a (z').OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z'"
using new by (simp add: fresh_fun_simp_OrL)
- also have "\<dots> \<longrightarrow>\<^isub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[z'\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z')[z'\<turnstile>n>y]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6054,32 +6054,32 @@
done
also have "\<dots> = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) y" using fs new
by (auto simp add: fresh_prod fresh_atm nrename_fresh subst_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) y"
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) y"
using ih1 ih2 by (auto intro: a_star_congs)
also have "\<dots> = (OrL (u).M (v).N z)[x\<turnstile>n>y]" using eq fs by simp
- finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" by simp
+ finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" by simp
next
case False
assume neq: "z\<noteq>x"
have "(OrL (u).M (v).N z){x:=<a>.Ax y a} = OrL (u).(M{x:=<a>.Ax y a}) (v).(N{x:=<a>.Ax y a}) z"
using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) z"
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[x\<turnstile>n>y]) (v).(N[x\<turnstile>n>y]) z"
using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" using fs neq by simp
+ finally show "(OrL (u).M (v).N z){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[x\<turnstile>n>y]" using fs neq by simp
qed
next
case (ImpR z c M d x a y)
have fs: "z\<sharp>x" "z\<sharp>a" "z\<sharp>y" "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "z\<sharp>d" "c\<sharp>d" by fact+
- have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
+ have ih: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
have "(ImpR (z).<c>.M d){x:=<a>.Ax y a} = ImpR (z).<c>.(M{x:=<a>.Ax y a}) d" using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
- finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[x\<turnstile>n>y]) d" using ih by (auto intro: a_star_congs)
+ finally show "(ImpR (z).<c>.M d){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[x\<turnstile>n>y]" using fs by simp
next
case (ImpL c M u N v x a y)
have fs: "c\<sharp>x" "c\<sharp>a" "c\<sharp>y" "u\<sharp>x" "u\<sharp>a" "u\<sharp>y" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
- have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* M[x\<turnstile>n>y]" by fact
- have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* N[x\<turnstile>n>y]" by fact
- show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
+ have ih1: "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]" by fact
+ have ih2: "N{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* N[x\<turnstile>n>y]" by fact
+ show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
proof(cases "v=x")
case True
assume eq: "v=x"
@@ -6090,7 +6090,7 @@
using eq fs by simp
also have "\<dots> = Cut <a>.Ax y a (v').ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v'"
using new by (simp add: fresh_fun_simp_ImpL)
- also have "\<dots> \<longrightarrow>\<^isub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v')[v'\<turnstile>n>y]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6100,65 +6100,65 @@
done
also have "\<dots> = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) y" using fs new
by (auto simp add: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) y"
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) y"
using ih1 ih2 by (auto intro: a_star_congs)
also have "\<dots> = (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using eq fs by simp
- finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using fs by simp
+ finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]" using fs by simp
next
case False
assume neq: "v\<noteq>x"
have "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} = ImpL <c>.(M{x:=<a>.Ax y a}) (u).(N{x:=<a>.Ax y a}) v"
using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) v"
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[x\<turnstile>n>y]) (u).(N[x\<turnstile>n>y]) v"
using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
+ finally show "(ImpL <c>.M (u).N v){x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[x\<turnstile>n>y]"
using fs neq by simp
qed
qed
lemma subst_with_ax2:
- shows "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]"
+ shows "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]"
proof(nominal_induct M avoiding: b a x rule: trm.strong_induct)
case (Ax z c b a x)
- show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]"
+ show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]"
proof (cases "c=b")
case True
assume eq: "c=b"
have "(Ax z c){b:=(x).Ax x a} = Cut <b>.Ax z c (x).Ax x a" using eq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" using eq by blast
- finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" using eq by blast
+ finally show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
next
case False
assume neq: "c\<noteq>b"
- then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Ax z c)[b\<turnstile>c>a]" by simp
+ then show "(Ax z c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Ax z c)[b\<turnstile>c>a]" by simp
qed
next
case (Cut c M z N b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>N" "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>M" by fact+
- have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
- show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]"
+ have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
+ show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]"
proof (cases "N = Ax z b")
case True
assume eq: "N = Ax z b"
have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (x).Ax x a" using eq fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs by blast
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (x).Ax x a" using ih1 a_star_congs by blast
also have "\<dots> = Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using eq fs
by (simp add: trm.inject alpha calc_atm fresh_atm)
- finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
+ finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
next
case False
assume neq: "N \<noteq> Ax z b"
have "(Cut <c>.M (z).N){b:=(x).Ax x a} = Cut <c>.(M{b:=(x).Ax x a}) (z).(N{b:=(x).Ax x a})"
using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs by blast
- finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <c>.(M[b\<turnstile>c>a]) (z).(N[b\<turnstile>c>a])" using ih1 ih2 a_star_congs by blast
+ finally show "(Cut <c>.M (z).N){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (Cut <c>.M (z).N)[b\<turnstile>c>a]" using fs by simp
qed
next
case (NotR z M c b a x)
have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "z\<sharp>c" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]"
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]"
proof (cases "c=b")
case True
assume eq: "c=b"
@@ -6168,7 +6168,7 @@
using eq fs by simp
also have "\<dots> = Cut <a'>.NotR (z).M{b:=(x).Ax x a} a' (x).Ax x a"
using new by (simp add: fresh_fun_simp_NotR fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\<turnstile>c>a]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6177,29 +6177,29 @@
apply(simp)
done
also have "\<dots> = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
also have "\<dots> = (NotR (z).M c)[b\<turnstile>c>a]" using eq by simp
- finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" by simp
+ finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" by simp
next
case False
assume neq: "c\<noteq>b"
have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih by (auto intro: a_star_congs)
- finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotR (z).M c)[b\<turnstile>c>a]" using neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotR (z).(M[b\<turnstile>c>a]) c" using ih by (auto intro: a_star_congs)
+ finally show "(NotR (z).M c){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotR (z).M c)[b\<turnstile>c>a]" using neq by simp
qed
next
case (NotL c M z b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>z" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
have "(NotL <c>.M z){b:=(x).Ax x a} = NotL <c>.(M{b:=(x).Ax x a}) z" using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih by (auto intro: a_star_congs)
- finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (NotL <c>.M z)[b\<turnstile>c>a]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* NotL <c>.(M[b\<turnstile>c>a]) z" using ih by (auto intro: a_star_congs)
+ finally show "(NotL <c>.M z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (NotL <c>.M z)[b\<turnstile>c>a]" using fs by simp
next
case (AndR c M d N e b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "d\<sharp>b" "d\<sharp>a" "d\<sharp>x" "d\<noteq>c" "c\<sharp>N" "c\<sharp>e" "d\<sharp>M" "d\<sharp>e" by fact+
- have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
- show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
+ have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
+ show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
proof(cases "e=b")
case True
assume eq: "e=b"
@@ -6210,7 +6210,7 @@
using eq fs by simp
also have "\<dots> = Cut <e'>.AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e' (x).Ax x a"
using new by (simp add: fresh_fun_simp_AndR fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e')[e'\<turnstile>c>a]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6220,37 +6220,37 @@
done
also have "\<dots> = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) a" using fs new
by (auto simp add: fresh_prod fresh_atm subst_fresh crename_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" using ih1 ih2 by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) a" using ih1 ih2 by (auto intro: a_star_congs)
also have "\<dots> = (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" using eq fs by simp
- finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" by simp
+ finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]" by simp
next
case False
assume neq: "e\<noteq>b"
have "(AndR <c>.M <d>.N e){b:=(x).Ax x a} = AndR <c>.(M{b:=(x).Ax x a}) <d>.(N{b:=(x).Ax x a}) e"
using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndR <c>.(M[b\<turnstile>c>a]) <d>.(N[b\<turnstile>c>a]) e" using ih1 ih2 by (auto intro: a_star_congs)
+ finally show "(AndR <c>.M <d>.N e){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndR <c>.M <d>.N e)[b\<turnstile>c>a]"
using fs neq by simp
qed
next
case (AndL1 u M v b a x)
have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
- finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL1 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
+ finally show "(AndL1 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL1 (u).M v)[b\<turnstile>c>a]" using fs by simp
next
case (AndL2 u M v b a x)
have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "u\<sharp>v" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
- finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* AndL2 (u).(M[b\<turnstile>c>a]) v" using ih by (auto intro: a_star_congs)
+ finally show "(AndL2 (u).M v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (AndL2 (u).M v)[b\<turnstile>c>a]" using fs by simp
next
case (OrR1 c M d b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]"
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]"
proof(cases "d=b")
case True
assume eq: "d=b"
@@ -6260,7 +6260,7 @@
fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
also have "\<dots> = Cut <a'>.OrR1 <c>.M{b:=(x).Ax x a} a' (x).Ax x a"
using new by (simp add: fresh_fun_simp_OrR1)
- also have "\<dots> \<longrightarrow>\<^isub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (OrR1 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6270,21 +6270,21 @@
done
also have "\<dots> = OrR1 <c>.M{b:=(x).Ax x a} a" using fs new
by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
also have "\<dots> = (OrR1 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
- finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" by simp
+ finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" by simp
next
case False
assume neq: "d\<noteq>b"
have "(OrR1 <c>.M d){b:=(x).Ax x a} = OrR1 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
- finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR1 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
+ finally show "(OrR1 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR1 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
qed
next
case (OrR2 c M d b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "c\<sharp>d" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]"
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]"
proof(cases "d=b")
case True
assume eq: "d=b"
@@ -6294,7 +6294,7 @@
fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp)
also have "\<dots> = Cut <a'>.OrR2 <c>.M{b:=(x).Ax x a} a' (x).Ax x a"
using new by (simp add: fresh_fun_simp_OrR2)
- also have "\<dots> \<longrightarrow>\<^isub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (OrR2 <c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6304,30 +6304,30 @@
done
also have "\<dots> = OrR2 <c>.M{b:=(x).Ax x a} a" using fs new
by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
also have "\<dots> = (OrR2 <c>.M d)[b\<turnstile>c>a]" using eq fs by simp
- finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" by simp
+ finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" by simp
next
case False
assume neq: "d\<noteq>b"
have "(OrR2 <c>.M d){b:=(x).Ax x a} = OrR2 <c>.(M{b:=(x).Ax x a}) d" using fs neq by (simp)
- also have "\<dots> \<longrightarrow>\<^isub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
- finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrR2 <c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
+ finally show "(OrR2 <c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrR2 <c>.M d)[b\<turnstile>c>a]" using fs neq by simp
qed
next
case (OrL u M v N z b a x)
have fs: "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "v\<sharp>b" "v\<sharp>a" "v\<sharp>x" "v\<noteq>u" "u\<sharp>N" "u\<sharp>z" "v\<sharp>M" "v\<sharp>z" by fact+
- have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
+ have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
have "(OrL (u).M (v).N z){b:=(x).Ax x a} = OrL (u).(M{b:=(x).Ax x a}) (v).(N{b:=(x).Ax x a}) z"
using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" using fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* OrL (u).(M[b\<turnstile>c>a]) (v).(N[b\<turnstile>c>a]) z" using ih1 ih2 by (auto intro: a_star_congs)
+ finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (OrL (u).M (v).N z)[b\<turnstile>c>a]" using fs by simp
next
case (ImpR z c M d b a x)
have fs: "z\<sharp>b" "z\<sharp>a" "z\<sharp>x" "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "z\<sharp>d" "c\<sharp>d" by fact+
- have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]"
+ have ih: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]"
proof(cases "b=d")
case True
assume eq: "b=d"
@@ -6337,7 +6337,7 @@
fresh_fun (\<lambda>a'. Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by simp
also have "\<dots> = Cut <a'>.ImpR z.<c>.M{b:=(x).Ax x a} a' (x).Ax x a"
using new by (simp add: fresh_fun_simp_ImpR)
- also have "\<dots> \<longrightarrow>\<^isub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (ImpR z.<c>.M{b:=(x).Ax x a} a')[a'\<turnstile>c>a]"
using new
apply(rule_tac a_starI)
apply(rule a_redu.intros)
@@ -6347,26 +6347,26 @@
done
also have "\<dots> = ImpR z.<c>.M{b:=(x).Ax x a} a" using fs new
by (auto simp add: fresh_prod crename_fresh subst_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpR z.<c>.(M[b\<turnstile>c>a]) a" using ih by (auto intro: a_star_congs)
also have "\<dots> = (ImpR z.<c>.M b)[b\<turnstile>c>a]" using eq fs by simp
- finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using eq by simp
+ finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using eq by simp
next
case False
assume neq: "b\<noteq>d"
have "(ImpR (z).<c>.M d){b:=(x).Ax x a} = ImpR (z).<c>.(M{b:=(x).Ax x a}) d" using fs neq by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
- finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq fs by simp
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpR (z).<c>.(M[b\<turnstile>c>a]) d" using ih by (auto intro: a_star_congs)
+ finally show "(ImpR (z).<c>.M d){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpR (z).<c>.M d)[b\<turnstile>c>a]" using neq fs by simp
qed
next
case (ImpL c M u N v b a x)
have fs: "c\<sharp>b" "c\<sharp>a" "c\<sharp>x" "u\<sharp>b" "u\<sharp>a" "u\<sharp>x" "c\<sharp>N" "c\<sharp>v" "u\<sharp>M" "u\<sharp>v" by fact+
- have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* M[b\<turnstile>c>a]" by fact
- have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^isub>a* N[b\<turnstile>c>a]" by fact
+ have ih1: "M{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* M[b\<turnstile>c>a]" by fact
+ have ih2: "N{b:=(x).Ax x a} \<longrightarrow>\<^sub>a* N[b\<turnstile>c>a]" by fact
have "(ImpL <c>.M (u).N v){b:=(x).Ax x a} = ImpL <c>.(M{b:=(x).Ax x a}) (u).(N{b:=(x).Ax x a}) v"
using fs by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>a]) v"
+ also have "\<dots> \<longrightarrow>\<^sub>a* ImpL <c>.(M[b\<turnstile>c>a]) (u).(N[b\<turnstile>c>a]) v"
using ih1 ih2 by (auto intro: a_star_congs)
- finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^isub>a* (ImpL <c>.M (u).N v)[b\<turnstile>c>a]"
+ finally show "(ImpL <c>.M (u).N v){b:=(x).Ax x a} \<longrightarrow>\<^sub>a* (ImpL <c>.M (u).N v)[b\<turnstile>c>a]"
using fs by simp
qed
--- a/src/HOL/Nominal/Examples/Class2.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Class2.thy Tue Aug 13 16:25:47 2013 +0200
@@ -27,8 +27,8 @@
qed
lemma c_redu_subst1:
- assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>M" "y\<sharp>P"
- shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
+ assumes a: "M \<longrightarrow>\<^sub>c M'" "c\<sharp>M" "y\<sharp>P"
+ shows "M{y:=<c>.P} \<longrightarrow>\<^sub>c M'{y:=<c>.P}"
using a
proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
case (left M a N x)
@@ -89,8 +89,8 @@
qed
lemma c_redu_subst2:
- assumes a: "M \<longrightarrow>\<^isub>c M'" "c\<sharp>P" "y\<sharp>M"
- shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
+ assumes a: "M \<longrightarrow>\<^sub>c M'" "c\<sharp>P" "y\<sharp>M"
+ shows "M{c:=(y).P} \<longrightarrow>\<^sub>c M'{c:=(y).P}"
using a
proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct)
case (right N x a M)
@@ -155,8 +155,8 @@
qed
lemma c_redu_subst1':
- assumes a: "M \<longrightarrow>\<^isub>c M'"
- shows "M{y:=<c>.P} \<longrightarrow>\<^isub>c M'{y:=<c>.P}"
+ assumes a: "M \<longrightarrow>\<^sub>c M'"
+ shows "M{y:=<c>.P} \<longrightarrow>\<^sub>c M'{y:=<c>.P}"
using a
proof -
obtain y'::"name" where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -169,7 +169,7 @@
apply(rule subst_rename(4))
apply(simp)
done
- also have "\<dots> \<longrightarrow>\<^isub>c ([(y',y)]\<bullet>M'){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
+ also have "\<dots> \<longrightarrow>\<^sub>c ([(y',y)]\<bullet>M'){y':=<c'>.([(c',c)]\<bullet>P)}" using fs1 fs2
apply -
apply(rule c_redu_subst1)
apply(simp add: c_redu.eqvt a)
@@ -188,8 +188,8 @@
qed
lemma c_redu_subst2':
- assumes a: "M \<longrightarrow>\<^isub>c M'"
- shows "M{c:=(y).P} \<longrightarrow>\<^isub>c M'{c:=(y).P}"
+ assumes a: "M \<longrightarrow>\<^sub>c M'"
+ shows "M{c:=(y).P} \<longrightarrow>\<^sub>c M'{c:=(y).P}"
using a
proof -
obtain y'::"name" where fs1: "y'\<sharp>(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast)
@@ -202,7 +202,7 @@
apply(rule subst_rename(2))
apply(simp)
done
- also have "\<dots> \<longrightarrow>\<^isub>c ([(c',c)]\<bullet>M'){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
+ also have "\<dots> \<longrightarrow>\<^sub>c ([(c',c)]\<bullet>M'){c':=(y').([(y',y)]\<bullet>P)}" using fs1 fs2
apply -
apply(rule c_redu_subst2)
apply(simp add: c_redu.eqvt a)
@@ -222,28 +222,28 @@
qed
lemma aux1:
- assumes a: "M = M'" "M' \<longrightarrow>\<^isub>l M''"
- shows "M \<longrightarrow>\<^isub>l M''"
+ assumes a: "M = M'" "M' \<longrightarrow>\<^sub>l M''"
+ shows "M \<longrightarrow>\<^sub>l M''"
using a by simp
lemma aux2:
- assumes a: "M \<longrightarrow>\<^isub>l M'" "M' = M''"
- shows "M \<longrightarrow>\<^isub>l M''"
+ assumes a: "M \<longrightarrow>\<^sub>l M'" "M' = M''"
+ shows "M \<longrightarrow>\<^sub>l M''"
using a by simp
lemma aux3:
- assumes a: "M = M'" "M' \<longrightarrow>\<^isub>a* M''"
- shows "M \<longrightarrow>\<^isub>a* M''"
+ assumes a: "M = M'" "M' \<longrightarrow>\<^sub>a* M''"
+ shows "M \<longrightarrow>\<^sub>a* M''"
using a by simp
lemma aux4:
assumes a: "M = M'"
- shows "M \<longrightarrow>\<^isub>a* M'"
+ shows "M \<longrightarrow>\<^sub>a* M'"
using a by blast
lemma l_redu_subst1:
- assumes a: "M \<longrightarrow>\<^isub>l M'"
- shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
+ assumes a: "M \<longrightarrow>\<^sub>l M'"
+ shows "M{y:=<c>.P} \<longrightarrow>\<^sub>a* M'{y:=<c>.P}"
using a
proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
case LAxR
@@ -302,7 +302,7 @@
have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} =
(Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using LNot
by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using LNot
+ also have "\<dots> \<longrightarrow>\<^sub>l (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using LNot
by (auto intro: l_redu.intros simp add: subst_fresh)
also have "\<dots> = (Cut <b>.N (u).M){y:=<c>.P}" using LNot asm
by (simp add: subst_fresh abs_fresh fresh_atm)
@@ -313,7 +313,7 @@
have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} =
(Cut <a>.NotR (u).(M{y:=<c>.P}) a (v).NotL <b>.(N{y:=<c>.P}) v)" using LNot
by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using LNot
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Cut <b>.(N{y:=<c>.P}) (u).(M{y:=<c>.P}))" using LNot
apply -
apply(rule a_starI)
apply(rule al_redu)
@@ -321,7 +321,7 @@
done
also have "\<dots> = (Cut <b>.(Cut <c>.P (y).Ax y b) (u).(M{y:=<c>.P}))" using LNot asm
by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(P[c\<turnstile>c>b]) (u).(M{y:=<c>.P}))"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Cut <b>.(P[c\<turnstile>c>b]) (u).(M{y:=<c>.P}))"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -356,7 +356,7 @@
apply(rule crename_swap)
apply(simp)
done
- finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){y:=<c>.P}"
+ finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <b>.N (u).M){y:=<c>.P}"
by simp
}
ultimately show ?thesis by blast
@@ -369,7 +369,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} =
Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z"
using LAnd1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
using LAnd1
apply -
apply(rule a_starI)
@@ -379,7 +379,7 @@
also have "\<dots> = (Cut <a1>.M1 (u).N){y:=<c>.P}" using LAnd1 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
by simp
}
moreover
@@ -387,7 +387,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} =
Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL1 (u).(N{y:=<c>.P}) z"
using LAnd1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.(M1{y:=<c>.P}) (u).(N{y:=<c>.P})"
using LAnd1
apply -
apply(rule a_starI)
@@ -396,7 +396,7 @@
done
also have "\<dots> = Cut <a1>.(Cut <c>.P (y). Ax y a1) (u).(N{y:=<c>.P})"
using LAnd1 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.P[c\<turnstile>c>a1] (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.P[c\<turnstile>c>a1] (u).(N{y:=<c>.P})"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -432,7 +432,7 @@
apply(simp)
done
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a1>.M1 (u).N){y:=<c>.P}"
by simp
}
ultimately show ?thesis by blast
@@ -445,7 +445,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} =
Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z"
using LAnd2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
using LAnd2
apply -
apply(rule a_starI)
@@ -455,7 +455,7 @@
also have "\<dots> = (Cut <a2>.M2 (u).N){y:=<c>.P}" using LAnd2 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
by simp
}
moreover
@@ -463,7 +463,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} =
Cut <b>.AndR <a1>.(M1{y:=<c>.P}) <a2>.(M2{y:=<c>.P}) b (z).AndL2 (u).(N{y:=<c>.P}) z"
using LAnd2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.(M2{y:=<c>.P}) (u).(N{y:=<c>.P})"
using LAnd2
apply -
apply(rule a_starI)
@@ -472,7 +472,7 @@
done
also have "\<dots> = Cut <a2>.(Cut <c>.P (y). Ax y a2) (u).(N{y:=<c>.P})"
using LAnd2 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.P[c\<turnstile>c>a2] (u).(N{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.P[c\<turnstile>c>a2] (u).(N{y:=<c>.P})"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -508,7 +508,7 @@
apply(simp)
done
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a2>.M2 (u).N){y:=<c>.P}"
by simp
}
ultimately show ?thesis by blast
@@ -521,7 +521,7 @@
have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} =
Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z"
using LOr1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
using LOr1
apply -
apply(rule a_starI)
@@ -531,7 +531,7 @@
also have "\<dots> = (Cut <a>.M (x1).N1){y:=<c>.P}" using LOr1 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
+ have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
by simp
}
moreover
@@ -539,7 +539,7 @@
have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} =
Cut <b>.OrR1 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z"
using LOr1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{y:=<c>.P}) (x1).(N1{y:=<c>.P})"
using LOr1
apply -
apply(rule a_starI)
@@ -548,7 +548,7 @@
done
also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x1).(N1{y:=<c>.P})"
using LOr1 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x1).(N1{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.P[c\<turnstile>c>a] (x1).(N1{y:=<c>.P})"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -584,7 +584,7 @@
apply(simp)
done
finally
- have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
+ have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x1).N1){y:=<c>.P}"
by simp
}
ultimately show ?thesis by blast
@@ -597,7 +597,7 @@
have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} =
Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z"
using LOr2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
using LOr2
apply -
apply(rule a_starI)
@@ -607,7 +607,7 @@
also have "\<dots> = (Cut <a>.M (x2).N2){y:=<c>.P}" using LOr2 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
+ have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
by simp
}
moreover
@@ -615,7 +615,7 @@
have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} =
Cut <b>.OrR2 <a>.(M{y:=<c>.P}) b (z).OrL (x1).(N1{y:=<c>.P}) (x2).(N2{y:=<c>.P}) z"
using LOr2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{y:=<c>.P}) (x2).(N2{y:=<c>.P})"
using LOr2
apply -
apply(rule a_starI)
@@ -624,7 +624,7 @@
done
also have "\<dots> = Cut <a>.(Cut <c>.P (y). Ax y a) (x2).(N2{y:=<c>.P})"
using LOr2 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.P[c\<turnstile>c>a] (x2).(N2{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.P[c\<turnstile>c>a] (x2).(N2{y:=<c>.P})"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -660,7 +660,7 @@
apply(simp)
done
finally
- have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
+ have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){y:=<c>.P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x2).N2){y:=<c>.P}"
by simp
}
ultimately show ?thesis by blast
@@ -673,7 +673,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} =
Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z"
using LImp by (simp add: fresh_prod abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P}) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P}) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
using LImp
apply -
apply(rule a_starI)
@@ -683,7 +683,7 @@
also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}" using LImp asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
by simp
}
@@ -692,7 +692,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} =
Cut <b>.ImpR (x).<a>.(M{y:=<c>.P}) b (z).ImpL <d>.(N{y:=<c>.P}) (u).(Q{y:=<c>.P}) z"
using LImp by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P}) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{y:=<c>.P}) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
using LImp
apply -
apply(rule a_starI)
@@ -701,7 +701,7 @@
done
also have "\<dots> = Cut <a>.(Cut <d>.(Cut <c>.P (y).Ax y d) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
using LImp asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(P[c\<turnstile>c>d]) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(P[c\<turnstile>c>d]) (x).(M{y:=<c>.P})) (u).(Q{y:=<c>.P})"
proof (cases "fic P c")
case True
assume "fic P c"
@@ -743,7 +743,7 @@
apply(simp)
done
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){y:=<c>.P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){y:=<c>.P}"
by simp
}
@@ -752,8 +752,8 @@
qed
lemma l_redu_subst2:
- assumes a: "M \<longrightarrow>\<^isub>l M'"
- shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
+ assumes a: "M \<longrightarrow>\<^sub>l M'"
+ shows "M{c:=(y).P} \<longrightarrow>\<^sub>a* M'{c:=(y).P}"
using a
proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct)
case LAxR
@@ -812,7 +812,7 @@
have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} =
(Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using LNot
by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>l (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using LNot
+ also have "\<dots> \<longrightarrow>\<^sub>l (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using LNot
by (auto intro: l_redu.intros simp add: subst_fresh)
also have "\<dots> = (Cut <b>.N (u).M){c:=(y).P}" using LNot asm
by (simp add: subst_fresh abs_fresh fresh_atm)
@@ -823,7 +823,7 @@
have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} =
(Cut <a>.NotR (u).(M{c:=(y).P}) a (v).NotL <b>.(N{c:=(y).P}) v)" using LNot
by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using LNot
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Cut <b>.(N{c:=(y).P}) (u).(M{c:=(y).P}))" using LNot
apply -
apply(rule a_starI)
apply(rule al_redu)
@@ -831,7 +831,7 @@
done
also have "\<dots> = (Cut <b>.(N{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P))" using LNot asm
by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* (Cut <b>.(N{c:=(y).P}) (u).(P[y\<turnstile>n>u]))"
+ also have "\<dots> \<longrightarrow>\<^sub>a* (Cut <b>.(N{c:=(y).P}) (u).(P[y\<turnstile>n>u]))"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -866,7 +866,7 @@
apply(rule nrename_swap)
apply(simp)
done
- finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <b>.N (u).M){c:=(y).P}"
+ finally have "(Cut <a>.NotR (u).M a (v).NotL <b>.N v){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <b>.N (u).M){c:=(y).P}"
by simp
}
ultimately show ?thesis by blast
@@ -879,7 +879,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} =
Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z"
using LAnd1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
using LAnd1
apply -
apply(rule a_starI)
@@ -889,7 +889,7 @@
also have "\<dots> = (Cut <a1>.M1 (u).N){c:=(y).P}" using LAnd1 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
by simp
}
moreover
@@ -897,7 +897,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} =
Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z"
using LAnd1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.(M1{c:=(y).P}) (u).(N{c:=(y).P})"
using LAnd1
apply -
apply(rule a_starI)
@@ -906,7 +906,7 @@
done
also have "\<dots> = Cut <a1>.(M1{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)"
using LAnd1 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a1>.(M1{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a1>.(M1{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -942,7 +942,7 @@
apply(simp)
done
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL1 (u).N z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a1>.M1 (u).N){c:=(y).P}"
by simp
}
ultimately show ?thesis by blast
@@ -955,7 +955,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} =
Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z"
using LAnd2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
using LAnd2
apply -
apply(rule a_starI)
@@ -965,7 +965,7 @@
also have "\<dots> = (Cut <a2>.M2 (u).N){c:=(y).P}" using LAnd2 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
by simp
}
moreover
@@ -973,7 +973,7 @@
have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} =
Cut <b>.AndR <a1>.(M1{c:=(y).P}) <a2>.(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z"
using LAnd2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.(M2{c:=(y).P}) (u).(N{c:=(y).P})"
using LAnd2
apply -
apply(rule a_starI)
@@ -982,7 +982,7 @@
done
also have "\<dots> = Cut <a2>.(M2{c:=(y).P}) (u).(Cut <c>.(Ax u c) (y).P)"
using LAnd2 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a2>.(M2{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a2>.(M2{c:=(y).P}) (u).(P[y\<turnstile>n>u])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1018,7 +1018,7 @@
apply(simp)
done
finally
- have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
+ have "(Cut <b>.AndR <a1>.M1 <a2>.M2 b (z).AndL2 (u).N z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a2>.M2 (u).N){c:=(y).P}"
by simp
}
ultimately show ?thesis by blast
@@ -1031,7 +1031,7 @@
have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} =
Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z"
using LOr1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
using LOr1
apply -
apply(rule a_starI)
@@ -1041,7 +1041,7 @@
also have "\<dots> = (Cut <a>.M (x1).N1){c:=(y).P}" using LOr1 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
+ have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
by simp
}
moreover
@@ -1049,7 +1049,7 @@
have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} =
Cut <b>.OrR1 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z"
using LOr1 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x1).(N1{c:=(y).P})"
using LOr1
apply -
apply(rule a_starI)
@@ -1058,7 +1058,7 @@
done
also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x1).(Cut <c>.(Ax x1 c) (y).P)"
using LOr1 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x1).(P[y\<turnstile>n>x1])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x1).(P[y\<turnstile>n>x1])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1094,7 +1094,7 @@
apply(simp)
done
finally
- have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
+ have "(Cut <b>.OrR1 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x1).N1){c:=(y).P}"
by simp
}
ultimately show ?thesis by blast
@@ -1107,7 +1107,7 @@
have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} =
Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z"
using LOr2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
using LOr2
apply -
apply(rule a_starI)
@@ -1117,7 +1117,7 @@
also have "\<dots> = (Cut <a>.M (x2).N2){c:=(y).P}" using LOr2 asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
+ have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
by simp
}
moreover
@@ -1125,7 +1125,7 @@
have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} =
Cut <b>.OrR2 <a>.(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z"
using LOr2 by (simp add: subst_fresh abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x2).(N2{c:=(y).P})"
using LOr2
apply -
apply(rule a_starI)
@@ -1134,7 +1134,7 @@
done
also have "\<dots> = Cut <a>.(M{c:=(y).P}) (x2).(Cut <c>.(Ax x2 c) (y).P)"
using LOr2 asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(M{c:=(y).P}) (x2).(P[y\<turnstile>n>x2])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(M{c:=(y).P}) (x2).(P[y\<turnstile>n>x2])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1170,7 +1170,7 @@
apply(simp)
done
finally
- have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^isub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
+ have "(Cut <b>.OrR2 <a>.M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \<longrightarrow>\<^sub>a* (Cut <a>.M (x2).N2){c:=(y).P}"
by simp
}
ultimately show ?thesis by blast
@@ -1183,7 +1183,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} =
Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z"
using LImp by (simp add: fresh_prod abs_fresh fresh_atm)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
using LImp
apply -
apply(rule a_starI)
@@ -1193,7 +1193,7 @@
also have "\<dots> = (Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}" using LImp asm
by (simp add: subst_fresh abs_fresh fresh_atm)
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
by simp
}
@@ -1202,7 +1202,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} =
Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z"
using LImp by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
using LImp
apply -
apply(rule a_starI)
@@ -1211,7 +1211,7 @@
done
also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(Cut <c>.Ax x c (y).P)) (u).(Q{c:=(y).P})"
using LImp asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(P[y\<turnstile>n>x])) (u).(Q{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(P[y\<turnstile>n>x])) (u).(Q{c:=(y).P})"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1251,7 +1251,7 @@
apply(simp add: nrename_swap)
done
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
by simp
}
@@ -1260,7 +1260,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} =
Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z"
using LImp by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
using LImp
apply -
apply(rule a_starI)
@@ -1269,7 +1269,7 @@
done
also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Cut <c>.Ax u c (y).P)"
using LImp asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(P[y\<turnstile>n>u])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(P[y\<turnstile>n>u])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1303,7 +1303,7 @@
apply(simp add: nrename_swap)
done
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
by simp
}
@@ -1312,7 +1312,7 @@
have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} =
Cut <b>.ImpR (x).<a>.(M{c:=(y).P}) b (z).ImpL <d>.(N{c:=(y).P}) (u).(Q{c:=(y).P}) z"
using LImp by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod)
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})"
using LImp
apply -
apply(rule a_starI)
@@ -1321,7 +1321,7 @@
done
also have "\<dots> = Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(Cut <c>.Ax x c (y).P)) (u).(Cut <c>.Ax u c (y).P)"
using LImp asm by simp
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(Cut <c>.Ax x c (y).P)) (u).(P[y\<turnstile>n>u])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(Cut <c>.Ax x c (y).P)) (u).(P[y\<turnstile>n>u])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1347,7 +1347,7 @@
apply(simp add: subst_with_ax1)
done
qed
- also have "\<dots> \<longrightarrow>\<^isub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(P[y\<turnstile>n>x])) (u).(P[y\<turnstile>n>u])"
+ also have "\<dots> \<longrightarrow>\<^sub>a* Cut <a>.(Cut <d>.(N{c:=(y).P}) (x).(P[y\<turnstile>n>x])) (u).(P[y\<turnstile>n>u])"
proof (cases "fin P y")
case True
assume "fin P y"
@@ -1386,7 +1386,7 @@
apply(simp add: nrename_swap)
done
finally
- have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^isub>a*
+ have "(Cut <b>.ImpR (x).<a>.M b (z).ImpL <d>.N (u).Q z){c:=(y).P} \<longrightarrow>\<^sub>a*
(Cut <a>.(Cut <d>.N (x).M) (u).Q){c:=(y).P}"
by simp
}
@@ -1395,8 +1395,8 @@
qed
lemma a_redu_subst1:
- assumes a: "M \<longrightarrow>\<^isub>a M'"
- shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
+ assumes a: "M \<longrightarrow>\<^sub>a M'"
+ shows "M{y:=<c>.P} \<longrightarrow>\<^sub>a* M'{y:=<c>.P}"
using a
proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
case al_redu
@@ -1602,8 +1602,8 @@
qed
lemma a_redu_subst2:
- assumes a: "M \<longrightarrow>\<^isub>a M'"
- shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
+ assumes a: "M \<longrightarrow>\<^sub>a M'"
+ shows "M{c:=(y).P} \<longrightarrow>\<^sub>a* M'{c:=(y).P}"
using a
proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct)
case al_redu
@@ -1802,8 +1802,8 @@
qed
lemma a_star_subst1:
- assumes a: "M \<longrightarrow>\<^isub>a* M'"
- shows "M{y:=<c>.P} \<longrightarrow>\<^isub>a* M'{y:=<c>.P}"
+ assumes a: "M \<longrightarrow>\<^sub>a* M'"
+ shows "M{y:=<c>.P} \<longrightarrow>\<^sub>a* M'{y:=<c>.P}"
using a
apply(induct)
apply(blast)
@@ -1812,8 +1812,8 @@
done
lemma a_star_subst2:
- assumes a: "M \<longrightarrow>\<^isub>a* M'"
- shows "M{c:=(y).P} \<longrightarrow>\<^isub>a* M'{c:=(y).P}"
+ assumes a: "M \<longrightarrow>\<^sub>a* M'"
+ shows "M{c:=(y).P} \<longrightarrow>\<^sub>a* M'{c:=(y).P}"
using a
apply(induct)
apply(blast)
@@ -1828,11 +1828,11 @@
inductive
SNa :: "trm \<Rightarrow> bool"
where
- SNaI: "(\<And>M'. M \<longrightarrow>\<^isub>a M' \<Longrightarrow> SNa M') \<Longrightarrow> SNa M"
+ SNaI: "(\<And>M'. M \<longrightarrow>\<^sub>a M' \<Longrightarrow> SNa M') \<Longrightarrow> SNa M"
lemma SNa_induct[consumes 1]:
assumes major: "SNa M"
- assumes hyp: "\<And>M'. SNa M' \<Longrightarrow> (\<forall>M''. M'\<longrightarrow>\<^isub>a M'' \<longrightarrow> P M'' \<Longrightarrow> P M')"
+ assumes hyp: "\<And>M'. SNa M' \<Longrightarrow> (\<forall>M''. M'\<longrightarrow>\<^sub>a M'' \<longrightarrow> P M'' \<Longrightarrow> P M')"
shows "P M"
apply (rule major[THEN SNa.induct])
apply (rule hyp)
@@ -1845,10 +1845,10 @@
assumes a_SNa: "SNa a"
and b_SNa: "SNa b"
and hyp: "\<And>x z.
- (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> SNa y) \<Longrightarrow>
- (\<And>y. x\<longrightarrow>\<^isub>a y \<Longrightarrow> P y z) \<Longrightarrow>
- (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> SNa u) \<Longrightarrow>
- (\<And>u. z\<longrightarrow>\<^isub>a u \<Longrightarrow> P x u) \<Longrightarrow> P x z"
+ (\<And>y. x\<longrightarrow>\<^sub>a y \<Longrightarrow> SNa y) \<Longrightarrow>
+ (\<And>y. x\<longrightarrow>\<^sub>a y \<Longrightarrow> P y z) \<Longrightarrow>
+ (\<And>u. z\<longrightarrow>\<^sub>a u \<Longrightarrow> SNa u) \<Longrightarrow>
+ (\<And>u. z\<longrightarrow>\<^sub>a u \<Longrightarrow> P x u) \<Longrightarrow> P x z"
shows "P a b"
proof -
from a_SNa
@@ -1873,20 +1873,20 @@
qed
lemma double_SNa:
- "\<lbrakk>SNa a; SNa b; \<forall>x z. ((\<forall>y. x\<longrightarrow>\<^isub>ay \<longrightarrow> P y z) \<and> (\<forall>u. z\<longrightarrow>\<^isub>a u \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
+ "\<lbrakk>SNa a; SNa b; \<forall>x z. ((\<forall>y. x\<longrightarrow>\<^sub>ay \<longrightarrow> P y z) \<and> (\<forall>u. z\<longrightarrow>\<^sub>a u \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
apply(rule_tac double_SNa_aux)
apply(assumption)+
apply(blast)
done
lemma a_preserves_SNa:
- assumes a: "SNa M" "M\<longrightarrow>\<^isub>a M'"
+ assumes a: "SNa M" "M\<longrightarrow>\<^sub>a M'"
shows "SNa M'"
using a
by (erule_tac SNa.cases) (simp)
lemma a_star_preserves_SNa:
- assumes a: "SNa M" and b: "M\<longrightarrow>\<^isub>a* M'"
+ assumes a: "SNa M" and b: "M\<longrightarrow>\<^sub>a* M'"
shows "SNa M'"
using b a
by (induct) (auto simp add: a_preserves_SNa)
@@ -5463,7 +5463,7 @@
with imp fs have "(x'):(M'{a':=(z).Ax z a'}) \<in> \<parallel>(A)\<parallel>" by (simp add: fresh_prod fresh_atm)
then have "SNa (M'{a':=(z).Ax z a'})" using ih2 by blast
moreover
- have "M'{a':=(z).Ax z a'} \<longrightarrow>\<^isub>a* M'[a'\<turnstile>c>a']" by (simp add: subst_with_ax2)
+ have "M'{a':=(z).Ax z a'} \<longrightarrow>\<^sub>a* M'[a'\<turnstile>c>a']" by (simp add: subst_with_ax2)
ultimately have "SNa (M'[a'\<turnstile>c>a'])" by (simp add: a_star_preserves_SNa)
then have "SNa M'" by (simp add: crename_id)
then have "SNa M" using eq by (simp add: ImpR_in_SNa)
@@ -5499,8 +5499,8 @@
text {* Main lemma 2 *}
lemma AXIOMS_preserved:
- shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> AXIOMSc B"
- and "(x):M \<in> AXIOMSn B \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> AXIOMSn B"
+ shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> <a>:M' \<in> AXIOMSc B"
+ and "(x):M \<in> AXIOMSn B \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> (x):M' \<in> AXIOMSn B"
apply(simp_all add: AXIOMSc_def AXIOMSn_def)
apply(auto simp add: ntrm.inject ctrm.inject alpha)
apply(drule ax_do_not_a_star_reduce)
@@ -5514,28 +5514,28 @@
done
lemma BINDING_preserved:
- shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
- and "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)"
+ shows "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> <a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
+ and "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>) \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> (x):M' \<in> BINDINGn B (\<parallel><B>\<parallel>)"
proof -
- assume red: "M \<longrightarrow>\<^isub>a* M'"
+ assume red: "M \<longrightarrow>\<^sub>a* M'"
assume asm: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
{
fix x::"name" and P::"trm"
from asm have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M{a:=(x).P})" by (simp add: BINDINGc_elim)
moreover
- have "M{a:=(x).P} \<longrightarrow>\<^isub>a* M'{a:=(x).P}" using red by (simp add: a_star_subst2)
+ have "M{a:=(x).P} \<longrightarrow>\<^sub>a* M'{a:=(x).P}" using red by (simp add: a_star_subst2)
ultimately
have "((x):P) \<in> (\<parallel>(B)\<parallel>) \<Longrightarrow> SNa (M'{a:=(x).P})" by (simp add: a_star_preserves_SNa)
}
then show "<a>:M' \<in> BINDINGc B (\<parallel>(B)\<parallel>)" by (auto simp add: BINDINGc_def)
next
- assume red: "M \<longrightarrow>\<^isub>a* M'"
+ assume red: "M \<longrightarrow>\<^sub>a* M'"
assume asm: "(x):M \<in> BINDINGn B (\<parallel><B>\<parallel>)"
{
fix c::"coname" and P::"trm"
from asm have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M{x:=<c>.P})" by (simp add: BINDINGn_elim)
moreover
- have "M{x:=<c>.P} \<longrightarrow>\<^isub>a* M'{x:=<c>.P}" using red by (simp add: a_star_subst1)
+ have "M{x:=<c>.P} \<longrightarrow>\<^sub>a* M'{x:=<c>.P}" using red by (simp add: a_star_subst1)
ultimately
have "(<c>:P) \<in> (\<parallel><B>\<parallel>) \<Longrightarrow> SNa (M'{x:=<c>.P})" by (simp add: a_star_preserves_SNa)
}
@@ -5543,12 +5543,12 @@
qed
lemma CANDs_preserved:
- shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
- and "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a* M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
+ shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
+ and "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
proof(nominal_induct B arbitrary: a x M M' rule: ty.strong_induct)
case (PR X)
{ case 1
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "<a>:M \<in> \<parallel><PR X>\<parallel>" by fact
then have "<a>:M \<in> NEGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
then have "<a>:M \<in> AXIOMSc (PR X) \<union> BINDINGc (PR X) (\<parallel>(PR X)\<parallel>)" by simp
@@ -5565,7 +5565,7 @@
then show "<a>:M' \<in> (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
next
case 2
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "(x):M \<in> \<parallel>(PR X)\<parallel>" by fact
then have "(x):M \<in> NEGn (PR X) (\<parallel><PR X>\<parallel>)" using NEG_simp by blast
then have "(x):M \<in> AXIOMSn (PR X) \<union> BINDINGn (PR X) (\<parallel><PR X>\<parallel>)" by simp
@@ -5583,12 +5583,12 @@
}
next
case (IMP A B)
- have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
- have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
- have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
- have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
+ have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
+ have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
+ have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
+ have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
{ case 1
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "<a>:M \<in> \<parallel><A IMP B>\<parallel>" by fact
then have "<a>:M \<in> NEGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)" by simp
then have "<a>:M \<in> AXIOMSc (A IMP B) \<union> BINDINGc (A IMP B) (\<parallel>(A IMP B)\<parallel>)
@@ -5607,7 +5607,7 @@
and imp1: "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N'{a':=(z).P}) \<in> \<parallel>(A)\<parallel>"
and imp2: "\<forall>c Q. a'\<sharp>(c,Q) \<and> <c>:Q \<in> \<parallel><A>\<parallel> \<longrightarrow> <a'>:(N'{x':=<c>.Q}) \<in> \<parallel><B>\<parallel>"
using IMPRIGHT_elim by blast
- from eq asm obtain N'' where eq': "M' = ImpR (x').<a'>.N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = ImpR (x').<a'>.N'' a" and red: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_ImpR_elim by (blast)
from imp1 have "\<forall>z P. x'\<sharp>(z,P) \<and> (z):P \<in> \<parallel>(B)\<parallel> \<longrightarrow> (x'):(N''{a':=(z).P}) \<in> \<parallel>(A)\<parallel>" using red ih2
apply(auto)
@@ -5636,7 +5636,7 @@
then show "<a>:M' \<in> (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
next
case 2
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "(x):M \<in> \<parallel>(A IMP B)\<parallel>" by fact
then have "(x):M \<in> NEGn (A IMP B) (\<parallel><A IMP B>\<parallel>)" using NEG_simp by blast
then have "(x):M \<in> AXIOMSn (A IMP B) \<union> BINDINGn (A IMP B) (\<parallel><A IMP B>\<parallel>)
@@ -5656,7 +5656,7 @@
and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "(y'):N' \<in> \<parallel>(B)\<parallel>"
by (erule_tac IMPLEFT_elim, blast)
from eq asm obtain T'' N'' where eq': "M' = ImpL <a'>.T'' (y').N'' x"
- and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
+ and red1: "T' \<longrightarrow>\<^sub>a* T''" and red2: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_ImpL_elim by blast
from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
moreover
@@ -5672,12 +5672,12 @@
}
next
case (AND A B)
- have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
- have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
- have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
- have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
+ have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
+ have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
+ have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
+ have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
{ case 1
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "<a>:M \<in> \<parallel><A AND B>\<parallel>" by fact
then have "<a>:M \<in> NEGc (A AND B) (\<parallel>(A AND B)\<parallel>)" by simp
then have "<a>:M \<in> AXIOMSc (A AND B) \<union> BINDINGc (A AND B) (\<parallel>(A AND B)\<parallel>)
@@ -5697,7 +5697,7 @@
and imp1: "<a'>:T' \<in> \<parallel><A>\<parallel>" and imp2: "<b'>:N' \<in> \<parallel><B>\<parallel>"
using ANDRIGHT_elim by blast
from eq asm obtain T'' N'' where eq': "M' = AndR <a'>.T'' <b'>.N'' a"
- and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
+ and red1: "T' \<longrightarrow>\<^sub>a* T''" and red2: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_AndR_elim by blast
from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
moreover
@@ -5712,7 +5712,7 @@
then show "<a>:M' \<in> (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
next
case 2
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "(x):M \<in> \<parallel>(A AND B)\<parallel>" by fact
then have "(x):M \<in> NEGn (A AND B) (\<parallel><A AND B>\<parallel>)" using NEG_simp by blast
then have "(x):M \<in> AXIOMSn (A AND B) \<union> BINDINGn (A AND B) (\<parallel><A AND B>\<parallel>)
@@ -5730,7 +5730,7 @@
then obtain y' N' where eq: "M = AndL1 (y').N' x"
and fin: "fin (AndL1 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
by (erule_tac ANDLEFT1_elim, blast)
- from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_AndL1_elim by blast
from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
moreover
@@ -5742,7 +5742,7 @@
then obtain y' N' where eq: "M = AndL2 (y').N' x"
and fin: "fin (AndL2 (y').N' x) x" and imp: "(y'):N' \<in> \<parallel>(B)\<parallel>"
by (erule_tac ANDLEFT2_elim, blast)
- from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_AndL2_elim by blast
from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
moreover
@@ -5756,12 +5756,12 @@
}
next
case (OR A B)
- have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
- have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
- have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
- have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
+ have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
+ have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
+ have ih3: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><B>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>" by fact
+ have ih4: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(B)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>" by fact
{ case 1
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "<a>:M \<in> \<parallel><A OR B>\<parallel>" by fact
then have "<a>:M \<in> NEGc (A OR B) (\<parallel>(A OR B)\<parallel>)" by simp
then have "<a>:M \<in> AXIOMSc (A OR B) \<union> BINDINGc (A OR B) (\<parallel>(A OR B)\<parallel>)
@@ -5779,7 +5779,7 @@
then obtain a' N' where eq: "M = OrR1 <a'>.N' a"
and fic: "fic (OrR1 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><A>\<parallel>"
using ORRIGHT1_elim by blast
- from eq asm obtain N'' where eq': "M' = OrR1 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = OrR1 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_OrR1_elim by blast
from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
moreover
@@ -5791,7 +5791,7 @@
then obtain a' N' where eq: "M = OrR2 <a'>.N' a"
and fic: "fic (OrR2 <a'>.N' a) a" and imp1: "<a'>:N' \<in> \<parallel><B>\<parallel>"
using ORRIGHT2_elim by blast
- from eq asm obtain N'' where eq': "M' = OrR2 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = OrR2 <a'>.N'' a" and red1: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_OrR2_elim by blast
from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
moreover
@@ -5804,7 +5804,7 @@
then show "<a>:M' \<in> (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
next
case 2
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "(x):M \<in> \<parallel>(A OR B)\<parallel>" by fact
then have "(x):M \<in> NEGn (A OR B) (\<parallel><A OR B>\<parallel>)" using NEG_simp by blast
then have "(x):M \<in> AXIOMSn (A OR B) \<union> BINDINGn (A OR B) (\<parallel><A OR B>\<parallel>)
@@ -5824,7 +5824,7 @@
and imp1: "(y'):T' \<in> \<parallel>(A)\<parallel>" and imp2: "(z'):N' \<in> \<parallel>(B)\<parallel>"
by (erule_tac ORLEFT_elim, blast)
from eq asm obtain T'' N'' where eq': "M' = OrL (y').T'' (z').N'' x"
- and red1: "T' \<longrightarrow>\<^isub>a* T''" and red2: "N' \<longrightarrow>\<^isub>a* N''"
+ and red1: "T' \<longrightarrow>\<^sub>a* T''" and red2: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_OrL_elim by blast
from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
moreover
@@ -5840,10 +5840,10 @@
}
next
case (NOT A)
- have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
- have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^isub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
+ have ih1: "\<And>a M M'. \<lbrakk><a>:M \<in> \<parallel><A>\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> <a>:M' \<in> \<parallel><A>\<parallel>" by fact
+ have ih2: "\<And>x M M'. \<lbrakk>(x):M \<in> \<parallel>(A)\<parallel>; M \<longrightarrow>\<^sub>a* M'\<rbrakk> \<Longrightarrow> (x):M' \<in> \<parallel>(A)\<parallel>" by fact
{ case 1
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "<a>:M \<in> \<parallel><NOT A>\<parallel>" by fact
then have "<a>:M \<in> NEGc (NOT A) (\<parallel>(NOT A)\<parallel>)" by simp
then have "<a>:M \<in> AXIOMSc (NOT A) \<union> BINDINGc (NOT A) (\<parallel>(NOT A)\<parallel>)
@@ -5861,7 +5861,7 @@
then obtain y' N' where eq: "M = NotR (y').N' a"
and fic: "fic (NotR (y').N' a) a" and imp: "(y'):N' \<in> \<parallel>(A)\<parallel>"
using NOTRIGHT_elim by blast
- from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_NotR_elim by blast
from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce)
moreover
@@ -5874,7 +5874,7 @@
then show "<a>:M' \<in> (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
next
case 2
- have asm: "M \<longrightarrow>\<^isub>a* M'" by fact
+ have asm: "M \<longrightarrow>\<^sub>a* M'" by fact
have "(x):M \<in> \<parallel>(NOT A)\<parallel>" by fact
then have "(x):M \<in> NEGn (NOT A) (\<parallel><NOT A>\<parallel>)" using NEG_simp by blast
then have "(x):M \<in> AXIOMSn (NOT A) \<union> BINDINGn (NOT A) (\<parallel><NOT A>\<parallel>)
@@ -5892,7 +5892,7 @@
then obtain a' N' where eq: "M = NotL <a'>.N' x"
and fin: "fin (NotL <a'>.N' x) x" and imp: "<a'>:N' \<in> \<parallel><A>\<parallel>"
by (erule_tac NOTLEFT_elim, blast)
- from eq asm obtain N'' where eq': "M' = NotL <a'>.N'' x" and red1: "N' \<longrightarrow>\<^isub>a* N''"
+ from eq asm obtain N'' where eq': "M' = NotL <a'>.N'' x" and red1: "N' \<longrightarrow>\<^sub>a* N''"
using a_star_redu_NotL_elim by blast
from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce)
moreover
@@ -5907,8 +5907,8 @@
qed
lemma CANDs_preserved_single:
- shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
- and "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^isub>a M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
+ shows "<a>:M \<in> \<parallel><B>\<parallel> \<Longrightarrow> M \<longrightarrow>\<^sub>a M' \<Longrightarrow> <a>:M' \<in> \<parallel><B>\<parallel>"
+ and "(x):M \<in> \<parallel>(B)\<parallel> \<Longrightarrow> M \<longrightarrow>\<^sub>a M' \<Longrightarrow> (x):M' \<in> \<parallel>(B)\<parallel>"
by (auto simp add: a_starI CANDs_preserved)
lemma fic_CANDS:
--- a/src/HOL/Nominal/Examples/Class3.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Class3.thy Tue Aug 13 16:25:47 2013 +0200
@@ -5,11 +5,11 @@
text {* 3rd Main Lemma *}
lemma Cut_a_redu_elim:
- assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>a R"
- shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^isub>a M') \<or>
- (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^isub>a N') \<or>
- (Cut <a>.M (x).N \<longrightarrow>\<^isub>c R) \<or>
- (Cut <a>.M (x).N \<longrightarrow>\<^isub>l R)"
+ assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
+ shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^sub>a M') \<or>
+ (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^sub>a N') \<or>
+ (Cut <a>.M (x).N \<longrightarrow>\<^sub>c R) \<or>
+ (Cut <a>.M (x).N \<longrightarrow>\<^sub>l R)"
using a
apply(erule_tac a_redu.cases)
apply(simp_all)
@@ -30,7 +30,7 @@
done
lemma Cut_c_redu_elim:
- assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^isub>c R"
+ assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>c R"
shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
(R = N{x:=<a>.M} \<and> \<not>fin N x)"
using a
@@ -563,8 +563,8 @@
done
lemma crename_credu:
- assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>c M'"
- shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>c M0"
+ assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>c M'"
+ shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>c M0"
using a
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
apply(drule sym)
@@ -594,8 +594,8 @@
done
lemma crename_lredu:
- assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>l M'"
- shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>l M0"
+ assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>l M'"
+ shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>l M0"
using a
apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
apply(drule sym)
@@ -805,8 +805,8 @@
done
lemma crename_aredu:
- assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^isub>a M'" "a\<noteq>b"
- shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^isub>a M0"
+ assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>a M'" "a\<noteq>b"
+ shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>a M0"
using a
apply(nominal_induct "M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
apply(drule crename_lredu)
@@ -1608,8 +1608,8 @@
done
lemma nrename_credu:
- assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>c M'"
- shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>c M0"
+ assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>c M'"
+ shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>c M0"
using a
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
apply(drule sym)
@@ -1658,8 +1658,8 @@
done
lemma nrename_lredu:
- assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>l M'"
- shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>l M0"
+ assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>l M'"
+ shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>l M0"
using a
apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
apply(drule sym)
@@ -1866,8 +1866,8 @@
done
lemma nrename_aredu:
- assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^isub>a M'" "x\<noteq>y"
- shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^isub>a M0"
+ assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>a M'" "x\<noteq>y"
+ shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>a M0"
using a
apply(nominal_induct "M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
apply(drule nrename_lredu)
@@ -2256,11 +2256,11 @@
abbreviation
A_Redu_set :: "(trm\<times>trm) set"
where
- "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^isub>a N}"
+ "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^sub>a N}"
lemma SNa_elim:
assumes a: "SNa M"
- shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^isub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
+ shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^sub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
using a
by (induct rule: SNa.induct) (blast)
@@ -2510,7 +2510,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x)
= Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
apply(simp)
-apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^isub>a Cut <b>.N' (y).M'a")
+apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
apply(simp only: a_preserves_SNa)
apply(rule al_redu)
apply(rule better_LNot_intro)
@@ -2531,7 +2531,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x')
= Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
apply(simp)
-apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^isub>a Cut <b>.N' (y).M'a")
+apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
apply(simp only: a_preserves_SNa)
apply(rule al_redu)
apply(rule better_LNot_intro)
@@ -2591,7 +2591,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x)
= Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
apply(simp)
-apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^isub>a Cut <b>.M1 (y).N'")
+apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LAnd1_intro)
@@ -2612,7 +2612,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z')
= Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
apply(simp)
-apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^isub>a Cut <b>.M1 (y).N'")
+apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LAnd1_intro)
@@ -2672,7 +2672,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x)
= Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
apply(simp)
-apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^isub>a Cut <c>.M2 (y).N'")
+apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LAnd2_intro)
@@ -2693,7 +2693,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z')
= Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
apply(simp)
-apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^isub>a Cut <c>.M2 (y).N'")
+apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LAnd2_intro)
@@ -2753,7 +2753,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x)
= Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
apply(simp)
-apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^isub>a Cut <b>.N' (z).M1")
+apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LOr1_intro)
@@ -2774,7 +2774,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z')
= Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
apply(simp)
-apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^isub>a Cut <b>.N' (z).M1")
+apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LOr1_intro)
@@ -2834,7 +2834,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x)
= Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
apply(simp)
-apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^isub>a Cut <b>.N' (y).M2")
+apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LOr2_intro)
@@ -2855,7 +2855,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z')
= Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
apply(simp)
-apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^isub>a Cut <b>.N' (y).M2")
+apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
apply(rule better_LOr2_intro)
@@ -2914,7 +2914,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
= Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
apply(simp)
-apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^isub>a
+apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^sub>a
Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
@@ -2937,7 +2937,7 @@
apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
= Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
apply(simp)
-apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^isub>a
+apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^sub>a
Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
apply(auto intro: a_preserves_SNa)[1]
apply(rule al_redu)
@@ -6097,7 +6097,7 @@
done
lemma id_redu:
- shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^isub>a* M"
+ shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^sub>a* M"
apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.strong_induct)
apply(auto)
(* Ax *)
@@ -6495,7 +6495,7 @@
fix a have "(idn \<Gamma> a) ncloses \<Gamma>" by (simp add: ncloses_id)
ultimately have "SNa ((idn \<Gamma> a),(idc \<Delta> x)<M>)" using a by (simp add: all_CAND)
moreover
- have "((idn \<Gamma> a),(idc \<Delta> x)<M>) \<longrightarrow>\<^isub>a* M" by (simp add: id_redu)
+ have "((idn \<Gamma> a),(idc \<Delta> x)<M>) \<longrightarrow>\<^sub>a* M" by (simp add: id_redu)
ultimately show "SNa M" by (simp add: a_star_preserves_SNa)
qed
--- a/src/HOL/Nominal/Examples/Contexts.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Contexts.thy Tue Aug 13 16:25:47 2013 +0200
@@ -44,7 +44,7 @@
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
where
"(Var x)[y::=s] = (if x=y then s else (Var x))"
-| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
+| "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])"
| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
apply(finite_guess)+
apply(rule TrueI)+
--- a/src/HOL/Nominal/Examples/Crary.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Crary.thy Tue Aug 13 16:25:47 2013 +0200
@@ -43,7 +43,7 @@
lemma ty_cases:
fixes T::ty
- shows "(\<exists> T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2) \<or> T=TUnit \<or> T=TBase"
+ shows "(\<exists> T\<^sub>1 T\<^sub>2. T=T\<^sub>1\<rightarrow>T\<^sub>2) \<or> T=TUnit \<or> T=TBase"
by (induct T rule:ty.induct) (auto)
instantiation ty :: size
@@ -53,7 +53,7 @@
where
"size (TBase) = 1"
| "size (TUnit) = 1"
-| "size (T\<^isub>1\<rightarrow>T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2"
+| "size (T\<^sub>1\<rightarrow>T\<^sub>2) = size T\<^sub>1 + size T\<^sub>2"
by (rule TrueI)+
instance ..
@@ -97,7 +97,7 @@
psubst :: "Subst \<Rightarrow> trm \<Rightarrow> trm" ("_<_>" [100,100] 130)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
-| "\<theta><(App t\<^isub>1 t\<^isub>2)> = App \<theta><t\<^isub>1> \<theta><t\<^isub>2>"
+| "\<theta><(App t\<^sub>1 t\<^sub>2)> = App \<theta><t\<^sub>1> \<theta><t\<^sub>2>"
| "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
| "\<theta><(Const n)> = Const n"
| "\<theta><(Unit)> = Unit"
@@ -114,7 +114,7 @@
lemma subst[simp]:
shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
- and "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])"
+ and "(App t\<^sub>1 t\<^sub>2)[y::=t'] = App (t\<^sub>1[y::=t']) (t\<^sub>2[y::=t'])"
and "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
and "Const n[y::=t'] = Const n"
and "Unit [y::=t'] = Unit"
@@ -128,10 +128,10 @@
lemma subst_rename:
fixes c::"name"
- assumes a: "c\<sharp>t\<^isub>1"
- shows "t\<^isub>1[a::=t\<^isub>2] = ([(c,a)]\<bullet>t\<^isub>1)[c::=t\<^isub>2]"
+ assumes a: "c\<sharp>t\<^sub>1"
+ shows "t\<^sub>1[a::=t\<^sub>2] = ([(c,a)]\<bullet>t\<^sub>1)[c::=t\<^sub>2]"
using a
-apply(nominal_induct t\<^isub>1 avoiding: a c t\<^isub>2 rule: trm.strong_induct)
+apply(nominal_induct t\<^sub>1 avoiding: a c t\<^sub>2 rule: trm.strong_induct)
apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+
done
@@ -145,24 +145,24 @@
lemma fresh_subst'':
fixes z::"name"
- assumes "z\<sharp>t\<^isub>2"
- shows "z\<sharp>t\<^isub>1[z::=t\<^isub>2]"
+ assumes "z\<sharp>t\<^sub>2"
+ shows "z\<sharp>t\<^sub>1[z::=t\<^sub>2]"
using assms
-by (nominal_induct t\<^isub>1 avoiding: t\<^isub>2 z rule: trm.strong_induct)
+by (nominal_induct t\<^sub>1 avoiding: t\<^sub>2 z rule: trm.strong_induct)
(auto simp add: abs_fresh fresh_nat fresh_atm)
lemma fresh_subst':
fixes z::"name"
- assumes "z\<sharp>[y].t\<^isub>1" "z\<sharp>t\<^isub>2"
- shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
+ assumes "z\<sharp>[y].t\<^sub>1" "z\<sharp>t\<^sub>2"
+ shows "z\<sharp>t\<^sub>1[y::=t\<^sub>2]"
using assms
-by (nominal_induct t\<^isub>1 avoiding: y t\<^isub>2 z rule: trm.strong_induct)
+by (nominal_induct t\<^sub>1 avoiding: y t\<^sub>2 z rule: trm.strong_induct)
(auto simp add: abs_fresh fresh_nat fresh_atm)
lemma fresh_subst:
fixes z::"name"
- assumes a: "z\<sharp>t\<^isub>1" "z\<sharp>t\<^isub>2"
- shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
+ assumes a: "z\<sharp>t\<^sub>1" "z\<sharp>t\<^sub>2"
+ shows "z\<sharp>t\<^sub>1[y::=t\<^sub>2]"
using a
by (auto simp add: fresh_subst' abs_fresh)
@@ -193,17 +193,17 @@
lemma subst_fun_eq:
fixes u::trm
- assumes h:"[x].t\<^isub>1 = [y].t\<^isub>2"
- shows "t\<^isub>1[x::=u] = t\<^isub>2[y::=u]"
+ assumes h:"[x].t\<^sub>1 = [y].t\<^sub>2"
+ shows "t\<^sub>1[x::=u] = t\<^sub>2[y::=u]"
proof -
{
- assume "x=y" and "t\<^isub>1=t\<^isub>2"
+ assume "x=y" and "t\<^sub>1=t\<^sub>2"
then have ?thesis using h by simp
}
moreover
{
- assume h1:"x \<noteq> y" and h2:"t\<^isub>1=[(x,y)] \<bullet> t\<^isub>2" and h3:"x \<sharp> t\<^isub>2"
- then have "([(x,y)] \<bullet> t\<^isub>2)[x::=u] = t\<^isub>2[y::=u]" by (simp add: subst_rename)
+ assume h1:"x \<noteq> y" and h2:"t\<^sub>1=[(x,y)] \<bullet> t\<^sub>2" and h3:"x \<sharp> t\<^sub>2"
+ then have "([(x,y)] \<bullet> t\<^sub>2)[x::=u] = t\<^sub>2[y::=u]" by (simp add: subst_rename)
then have ?thesis using h2 by simp
}
ultimately show ?thesis using alpha h by blast
@@ -275,13 +275,13 @@
abbreviation
"sub_context" :: "Ctxt \<Rightarrow> Ctxt \<Rightarrow> bool" (" _ \<subseteq> _ " [55,55] 55)
where
- "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>a T. (a,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (a,T)\<in>set \<Gamma>\<^isub>2"
+ "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>a T. (a,T)\<in>set \<Gamma>\<^sub>1 \<longrightarrow> (a,T)\<in>set \<Gamma>\<^sub>2"
lemma valid_monotonicity[elim]:
fixes \<Gamma> \<Gamma>' :: Ctxt
assumes a: "\<Gamma> \<subseteq> \<Gamma>'"
and b: "x\<sharp>\<Gamma>'"
- shows "(x,T\<^isub>1)#\<Gamma> \<subseteq> (x,T\<^isub>1)#\<Gamma>'"
+ shows "(x,T\<^sub>1)#\<Gamma> \<subseteq> (x,T\<^sub>1)#\<Gamma>'"
using a b by auto
lemma fresh_context:
@@ -295,9 +295,9 @@
lemma type_unicity_in_context:
assumes a: "valid \<Gamma>"
- and b: "(x,T\<^isub>1) \<in> set \<Gamma>"
- and c: "(x,T\<^isub>2) \<in> set \<Gamma>"
- shows "T\<^isub>1=T\<^isub>2"
+ and b: "(x,T\<^sub>1) \<in> set \<Gamma>"
+ and c: "(x,T\<^sub>2) \<in> set \<Gamma>"
+ shows "T\<^sub>1=T\<^sub>2"
using a b c
by (induct \<Gamma>)
(auto dest!: fresh_context)
@@ -306,8 +306,8 @@
typing :: "Ctxt\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
where
T_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| T_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> e\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2"
-| T_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2"
+| T_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> e\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^sub>1 e\<^sub>2 : T\<^sub>2"
+| T_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^sub>1\<rightarrow>T\<^sub>2"
| T_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : TBase"
| T_Unit[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Unit : TUnit"
@@ -344,12 +344,12 @@
Q_Refl[intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<equiv> t : T"
| Q_Symm[intro]: "\<Gamma> \<turnstile> t \<equiv> s : T \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"
| Q_Trans[intro]: "\<lbrakk>\<Gamma> \<turnstile> s \<equiv> t : T; \<Gamma> \<turnstile> t \<equiv> u : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> u : T"
-| Q_Abs[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s\<^isub>2 \<equiv> Lam [x]. t\<^isub>2 : T\<^isub>1 \<rightarrow> T\<^isub>2"
-| Q_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App s\<^isub>1 s\<^isub>2 \<equiv> App t\<^isub>1 t\<^isub>2 : T\<^isub>2"
-| Q_Beta[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s\<^isub>2,t\<^isub>2); (x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk>
- \<Longrightarrow> \<Gamma> \<turnstile> App (Lam [x]. s\<^isub>1) s\<^isub>2 \<equiv> t\<^isub>1[x::=t\<^isub>2] : T\<^isub>2"
-| Q_Ext[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2\<rbrakk>
- \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| Q_Abs[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^sub>1)#\<Gamma> \<turnstile> s\<^sub>2 \<equiv> t\<^sub>2 : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s\<^sub>2 \<equiv> Lam [x]. t\<^sub>2 : T\<^sub>1 \<rightarrow> T\<^sub>2"
+| Q_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> s\<^sub>1 \<equiv> t\<^sub>1 : T\<^sub>1 \<rightarrow> T\<^sub>2 ; \<Gamma> \<turnstile> s\<^sub>2 \<equiv> t\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App s\<^sub>1 s\<^sub>2 \<equiv> App t\<^sub>1 t\<^sub>2 : T\<^sub>2"
+| Q_Beta[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s\<^sub>2,t\<^sub>2); (x,T\<^sub>1)#\<Gamma> \<turnstile> s\<^sub>1 \<equiv> t\<^sub>1 : T\<^sub>2 ; \<Gamma> \<turnstile> s\<^sub>2 \<equiv> t\<^sub>2 : T\<^sub>1\<rbrakk>
+ \<Longrightarrow> \<Gamma> \<turnstile> App (Lam [x]. s\<^sub>1) s\<^sub>2 \<equiv> t\<^sub>1[x::=t\<^sub>2] : T\<^sub>2"
+| Q_Ext[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^sub>2\<rbrakk>
+ \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T\<^sub>1 \<rightarrow> T\<^sub>2"
| Q_Unit[intro]: "\<lbrakk>\<Gamma> \<turnstile> s : TUnit; \<Gamma> \<turnstile> t: TUnit\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : TUnit"
equivariance def_equiv
@@ -367,8 +367,8 @@
inductive
whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80)
where
- QAR_Beta[intro]: "App (Lam [x]. t\<^isub>1) t\<^isub>2 \<leadsto> t\<^isub>1[x::=t\<^isub>2]"
-| QAR_App[intro]: "t\<^isub>1 \<leadsto> t\<^isub>1' \<Longrightarrow> App t\<^isub>1 t\<^isub>2 \<leadsto> App t\<^isub>1' t\<^isub>2"
+ QAR_Beta[intro]: "App (Lam [x]. t\<^sub>1) t\<^sub>2 \<leadsto> t\<^sub>1[x::=t\<^sub>2]"
+| QAR_App[intro]: "t\<^sub>1 \<leadsto> t\<^sub>1' \<Longrightarrow> App t\<^sub>1 t\<^sub>2 \<leadsto> App t\<^sub>1' t\<^sub>2"
declare trm.inject [simp add]
declare ty.inject [simp add]
@@ -448,11 +448,11 @@
alg_path_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60,60,60] 60)
where
QAT_Base[intro]: "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
-| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2\<rbrakk>
- \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^sub>2\<rbrakk>
+ \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^sub>1 \<rightarrow> T\<^sub>2"
| QAT_One[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TUnit"
| QAP_Var[intro]: "\<lbrakk>valid \<Gamma>;(x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T"
-| QAP_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2"
+| QAP_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^sub>1 \<rightarrow> T\<^sub>2; \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T\<^sub>2"
| QAP_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n \<leftrightarrow> Const n : TBase"
equivariance alg_equiv
@@ -466,10 +466,10 @@
inductive_cases alg_equiv_inv_auto[elim]:
"\<Gamma> \<turnstile> s\<Leftrightarrow>t : TBase"
- "\<Gamma> \<turnstile> s\<Leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+ "\<Gamma> \<turnstile> s\<Leftrightarrow>t : T\<^sub>1 \<rightarrow> T\<^sub>2"
"\<Gamma> \<turnstile> s\<leftrightarrow>t : TBase"
"\<Gamma> \<turnstile> s\<leftrightarrow>t : TUnit"
- "\<Gamma> \<turnstile> s\<leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+ "\<Gamma> \<turnstile> s\<leftrightarrow>t : T\<^sub>1 \<rightarrow> T\<^sub>2"
"\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T"
"\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T'"
@@ -487,12 +487,12 @@
lemma Q_Arrow_strong_inversion:
assumes fs: "x\<sharp>\<Gamma>" "x\<sharp>t" "x\<sharp>u"
- and h: "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2"
- shows "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2"
+ and h: "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^sub>1\<rightarrow>T\<^sub>2"
+ shows "(x,T\<^sub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^sub>2"
proof -
- obtain y where fs2: "y\<sharp>(\<Gamma>,t,u)" and "(y,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T\<^isub>2"
+ obtain y where fs2: "y\<sharp>(\<Gamma>,t,u)" and "(y,T\<^sub>1)#\<Gamma> \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T\<^sub>2"
using h by auto
- then have "([(x,y)]\<bullet>((y,T\<^isub>1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) \<Leftrightarrow> [(x,y)]\<bullet> App u (Var y) : T\<^isub>2"
+ then have "([(x,y)]\<bullet>((y,T\<^sub>1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) \<Leftrightarrow> [(x,y)]\<bullet> App u (Var y) : T\<^sub>2"
using alg_equiv.eqvt[simplified] by blast
then show ?thesis using fs fs2 by (perm_simp)
qed
@@ -518,12 +518,12 @@
moreover have "valid \<Gamma>" "(x,T) \<in> set \<Gamma>" by fact+
ultimately show "T=T'" using type_unicity_in_context by auto
next
- case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u T\<^isub>2')
- have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T\<^isub>1\<rightarrow>T\<^isub>2 = T" by fact
- have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2'" by fact
- then obtain r t T\<^isub>1' where "u = App r t" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
- with ih have "T\<^isub>1\<rightarrow>T\<^isub>2 = T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
- then show "T\<^isub>2=T\<^isub>2'" using ty.inject by auto
+ case (QAP_App \<Gamma> p q T\<^sub>1 T\<^sub>2 s t u T\<^sub>2')
+ have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T\<^sub>1\<rightarrow>T\<^sub>2 = T" by fact
+ have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^sub>2'" by fact
+ then obtain r t T\<^sub>1' where "u = App r t" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^sub>1' \<rightarrow> T\<^sub>2'" by auto
+ with ih have "T\<^sub>1\<rightarrow>T\<^sub>2 = T\<^sub>1' \<rightarrow> T\<^sub>2'" by auto
+ then show "T\<^sub>2=T\<^sub>2'" using ty.inject by auto
qed (auto)
lemma alg_path_equiv_implies_valid:
@@ -552,28 +552,28 @@
have "s \<Down> p" by fact
ultimately show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : TBase" using b2 by auto
next
- case (QAT_Arrow x \<Gamma> s t T\<^isub>1 T\<^isub>2 u)
- have ih:"(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2
- \<Longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" by fact
+ case (QAT_Arrow x \<Gamma> s t T\<^sub>1 T\<^sub>2 u)
+ have ih:"(x,T\<^sub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^sub>2
+ \<Longrightarrow> (x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^sub>2" by fact
have fs: "x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" "x\<sharp>u" by fact+
- have "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
- then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" using fs
+ have "\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
+ then have "(x,T\<^sub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^sub>2" using fs
by (simp add: Q_Arrow_strong_inversion)
- with ih have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" by simp
- then show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
+ with ih have "(x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^sub>2" by simp
+ then show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : T\<^sub>1\<rightarrow>T\<^sub>2" using fs by (auto simp add: fresh_prod)
next
- case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u)
- have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T\<^isub>2" by fact
- then obtain r T\<^isub>1' v where ha: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^isub>1'\<rightarrow>T\<^isub>2" and hb: "\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^isub>1'" and eq: "u = App r v"
+ case (QAP_App \<Gamma> p q T\<^sub>1 T\<^sub>2 s t u)
+ have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T\<^sub>2" by fact
+ then obtain r T\<^sub>1' v where ha: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^sub>1'\<rightarrow>T\<^sub>2" and hb: "\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^sub>1'" and eq: "u = App r v"
by auto
- have ih1: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
- have ih2:"\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^isub>1" by fact
- have "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
- then have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T\<^isub>1\<rightarrow>T\<^isub>2" by (simp add: algorithmic_symmetry)
- with ha have "T\<^isub>1'\<rightarrow>T\<^isub>2 = T\<^isub>1\<rightarrow>T\<^isub>2" using algorithmic_path_type_unicity by simp
- then have "T\<^isub>1' = T\<^isub>1" by (simp add: ty.inject)
- then have "\<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^isub>1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2" using ih1 ih2 ha hb by auto
- then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2" using eq by auto
+ have ih1: "\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^sub>1\<rightarrow>T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
+ have ih2:"\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^sub>1 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^sub>1" by fact
+ have "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
+ then have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T\<^sub>1\<rightarrow>T\<^sub>2" by (simp add: algorithmic_symmetry)
+ with ha have "T\<^sub>1'\<rightarrow>T\<^sub>2 = T\<^sub>1\<rightarrow>T\<^sub>2" using algorithmic_path_type_unicity by simp
+ then have "T\<^sub>1' = T\<^sub>1" by (simp add: ty.inject)
+ then have "\<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^sub>1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^sub>1\<rightarrow>T\<^sub>2" using ih1 ih2 ha hb by auto
+ then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^sub>2" using eq by auto
qed (auto)
lemma algorithmic_weak_head_closure:
@@ -587,16 +587,16 @@
shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<subseteq> \<Gamma>' \<Longrightarrow> valid \<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<Leftrightarrow> t : T"
and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<subseteq> \<Gamma>' \<Longrightarrow> valid \<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv.strong_inducts)
- case (QAT_Arrow x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
+ case (QAT_Arrow x \<Gamma> s t T\<^sub>1 T\<^sub>2 \<Gamma>')
have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" "x\<sharp>\<Gamma>'" by fact+
have h2:"\<Gamma> \<subseteq> \<Gamma>'" by fact
- have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T\<^isub>1)#\<Gamma> \<subseteq> \<Gamma>'; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" by fact
+ have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T\<^sub>1)#\<Gamma> \<subseteq> \<Gamma>'; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^sub>2" by fact
have "valid \<Gamma>'" by fact
- then have "valid ((x,T\<^isub>1)#\<Gamma>')" using fs by auto
+ then have "valid ((x,T\<^sub>1)#\<Gamma>')" using fs by auto
moreover
- have sub: "(x,T\<^isub>1)#\<Gamma> \<subseteq> (x,T\<^isub>1)#\<Gamma>'" using h2 by auto
- ultimately have "(x,T\<^isub>1)#\<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih by simp
- then show "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
+ have sub: "(x,T\<^sub>1)#\<Gamma> \<subseteq> (x,T\<^sub>1)#\<Gamma>'" using h2 by auto
+ ultimately have "(x,T\<^sub>1)#\<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^sub>2" using ih by simp
+ then show "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^sub>1\<rightarrow>T\<^sub>2" using fs by (auto simp add: fresh_prod)
qed (auto)
lemma path_equiv_implies_nf:
@@ -611,10 +611,10 @@
where
"\<Gamma> \<turnstile> s is t : TUnit = True"
| "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
- | "\<Gamma> \<turnstile> s is t : (T\<^isub>1 \<rightarrow> T\<^isub>2) =
- (\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2))"
+ | "\<Gamma> \<turnstile> s is t : (T\<^sub>1 \<rightarrow> T\<^sub>2) =
+ (\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^sub>1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^sub>2))"
apply (auto simp add: ty.inject)
-apply (subgoal_tac "(\<exists>T\<^isub>1 T\<^isub>2. b=T\<^isub>1 \<rightarrow> T\<^isub>2) \<or> b=TUnit \<or> b=TBase" )
+apply (subgoal_tac "(\<exists>T\<^sub>1 T\<^sub>2. b=T\<^sub>1 \<rightarrow> T\<^sub>2) \<or> b=TUnit \<or> b=TBase" )
apply (force)
apply (rule ty_cases)
done
@@ -632,45 +632,45 @@
case (2 \<Gamma> s t \<Gamma>')
then show "\<Gamma>' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto
next
- case (3 \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
- have "\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2"
+ case (3 \<Gamma> s t T\<^sub>1 T\<^sub>2 \<Gamma>')
+ have "\<Gamma> \<turnstile> s is t : T\<^sub>1\<rightarrow>T\<^sub>2"
and "\<Gamma> \<subseteq> \<Gamma>'"
and "valid \<Gamma>'" by fact+
- then show "\<Gamma>' \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by simp
+ then show "\<Gamma>' \<turnstile> s is t : T\<^sub>1\<rightarrow>T\<^sub>2" by simp
qed (auto)
lemma main_lemma:
shows "\<Gamma> \<turnstile> s is t : T \<Longrightarrow> valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T"
and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> \<Gamma> \<turnstile> p is q : T"
proof (nominal_induct T arbitrary: \<Gamma> s t p q rule: ty.strong_induct)
- case (Arrow T\<^isub>1 T\<^isub>2)
+ case (Arrow T\<^sub>1 T\<^sub>2)
{
case (1 \<Gamma> s t)
- have ih1:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>2" by fact
- have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>1" by fact
- have h:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+ have ih1:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^sub>2; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^sub>2" by fact
+ have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^sub>1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^sub>1" by fact
+ have h:"\<Gamma> \<turnstile> s is t : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
obtain x::name where fs:"x\<sharp>(\<Gamma>,s,t)" by (erule exists_fresh[OF fs_name1])
have "valid \<Gamma>" by fact
- then have v: "valid ((x,T\<^isub>1)#\<Gamma>)" using fs by auto
- then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T\<^isub>1" by auto
- then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x is Var x : T\<^isub>1" using ih2 by auto
- then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T\<^isub>2" using h v by auto
- then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih1 v by auto
- then show "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
+ then have v: "valid ((x,T\<^sub>1)#\<Gamma>)" using fs by auto
+ then have "(x,T\<^sub>1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T\<^sub>1" by auto
+ then have "(x,T\<^sub>1)#\<Gamma> \<turnstile> Var x is Var x : T\<^sub>1" using ih2 by auto
+ then have "(x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T\<^sub>2" using h v by auto
+ then have "(x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^sub>2" using ih1 v by auto
+ then show "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^sub>1\<rightarrow>T\<^sub>2" using fs by (auto simp add: fresh_prod)
next
case (2 \<Gamma> p q)
- have h: "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
- have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>2" by fact
- have ih2:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>1; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" by fact
+ have h: "\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
+ have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^sub>2" by fact
+ have ih2:"\<And>\<Gamma> s t. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^sub>1; valid \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^sub>1" by fact
{
fix \<Gamma>' s t
- assume "\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T\<^isub>1" and hk: "valid \<Gamma>'"
- then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2" using h algorithmic_monotonicity by auto
- moreover have "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" using ih2 hl hk by auto
- ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2" by auto
- then have "\<Gamma>' \<turnstile> App p s is App q t : T\<^isub>2" using ih1 by auto
+ assume "\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T\<^sub>1" and hk: "valid \<Gamma>'"
+ then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T\<^sub>1 \<rightarrow> T\<^sub>2" using h algorithmic_monotonicity by auto
+ moreover have "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^sub>1" using ih2 hl hk by auto
+ ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T\<^sub>2" by auto
+ then have "\<Gamma>' \<turnstile> App p s is App q t : T\<^sub>2" using ih1 by auto
}
- then show "\<Gamma> \<turnstile> p is q : T\<^isub>1\<rightarrow>T\<^isub>2" by simp
+ then show "\<Gamma> \<turnstile> p is q : T\<^sub>1\<rightarrow>T\<^sub>2" by simp
}
next
case TBase
@@ -703,21 +703,21 @@
case TBase
then show "\<Gamma> \<turnstile> s is u : TBase" by (auto elim: algorithmic_transitivity)
next
- case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t u)
- have h1:"\<Gamma> \<turnstile> s is t : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
- have h2:"\<Gamma> \<turnstile> t is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
- have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>1; \<Gamma> \<turnstile> t is u : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>1" by fact
- have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; \<Gamma> \<turnstile> t is u : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>2" by fact
+ case (Arrow T\<^sub>1 T\<^sub>2 \<Gamma> s t u)
+ have h1:"\<Gamma> \<turnstile> s is t : T\<^sub>1 \<rightarrow> T\<^sub>2" by fact
+ have h2:"\<Gamma> \<turnstile> t is u : T\<^sub>1 \<rightarrow> T\<^sub>2" by fact
+ have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^sub>1; \<Gamma> \<turnstile> t is u : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^sub>1" by fact
+ have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^sub>2; \<Gamma> \<turnstile> t is u : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^sub>2" by fact
{
fix \<Gamma>' s' u'
- assume hsub:"\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T\<^isub>1" and hk: "valid \<Gamma>'"
- then have "\<Gamma>' \<turnstile> u' is s' : T\<^isub>1" using logical_symmetry by blast
- then have "\<Gamma>' \<turnstile> u' is u' : T\<^isub>1" using ih1 hl by blast
- then have "\<Gamma>' \<turnstile> App t u' is App u u' : T\<^isub>2" using h2 hsub hk by auto
- moreover have "\<Gamma>' \<turnstile> App s s' is App t u' : T\<^isub>2" using h1 hsub hl hk by auto
- ultimately have "\<Gamma>' \<turnstile> App s s' is App u u' : T\<^isub>2" using ih2 by blast
+ assume hsub:"\<Gamma> \<subseteq> \<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T\<^sub>1" and hk: "valid \<Gamma>'"
+ then have "\<Gamma>' \<turnstile> u' is s' : T\<^sub>1" using logical_symmetry by blast
+ then have "\<Gamma>' \<turnstile> u' is u' : T\<^sub>1" using ih1 hl by blast
+ then have "\<Gamma>' \<turnstile> App t u' is App u u' : T\<^sub>2" using h2 hsub hk by auto
+ moreover have "\<Gamma>' \<turnstile> App s s' is App t u' : T\<^sub>2" using h1 hsub hl hk by auto
+ ultimately have "\<Gamma>' \<turnstile> App s s' is App u u' : T\<^sub>2" using ih2 by blast
}
- then show "\<Gamma> \<turnstile> s is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
+ then show "\<Gamma> \<turnstile> s is u : T\<^sub>1 \<rightarrow> T\<^sub>2" by auto
qed (auto)
lemma logical_weak_head_closure:
@@ -740,21 +740,21 @@
case (TUnit \<Gamma> s t s')
then show ?case by auto
next
- case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t s')
+ case (Arrow T\<^sub>1 T\<^sub>2 \<Gamma> s t s')
have h1:"s' \<leadsto> s" by fact
- have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T\<^isub>2" by fact
- have h2:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+ have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^sub>2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T\<^sub>2" by fact
+ have h2:"\<Gamma> \<turnstile> s is t : T\<^sub>1\<rightarrow>T\<^sub>2" by fact
then
- have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2)"
+ have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<subseteq>\<Gamma>' \<longrightarrow> valid \<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^sub>1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^sub>2)"
by auto
{
- fix \<Gamma>' s\<^isub>2 t\<^isub>2
- assume "\<Gamma> \<subseteq> \<Gamma>'" and "\<Gamma>' \<turnstile> s\<^isub>2 is t\<^isub>2 : T\<^isub>1" and "valid \<Gamma>'"
- then have "\<Gamma>' \<turnstile> (App s s\<^isub>2) is (App t t\<^isub>2) : T\<^isub>2" using hb by auto
- moreover have "(App s' s\<^isub>2) \<leadsto> (App s s\<^isub>2)" using h1 by auto
- ultimately have "\<Gamma>' \<turnstile> App s' s\<^isub>2 is App t t\<^isub>2 : T\<^isub>2" using ih by auto
+ fix \<Gamma>' s\<^sub>2 t\<^sub>2
+ assume "\<Gamma> \<subseteq> \<Gamma>'" and "\<Gamma>' \<turnstile> s\<^sub>2 is t\<^sub>2 : T\<^sub>1" and "valid \<Gamma>'"
+ then have "\<Gamma>' \<turnstile> (App s s\<^sub>2) is (App t t\<^sub>2) : T\<^sub>2" using hb by auto
+ moreover have "(App s' s\<^sub>2) \<leadsto> (App s s\<^sub>2)" using h1 by auto
+ ultimately have "\<Gamma>' \<turnstile> App s' s\<^sub>2 is App t t\<^sub>2 : T\<^sub>2" using ih by auto
}
- then show "\<Gamma> \<turnstile> s' is t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
+ then show "\<Gamma> \<turnstile> s' is t : T\<^sub>1\<rightarrow>T\<^sub>2" by auto
qed
abbreviation
@@ -816,25 +816,25 @@
shows "\<Gamma>' \<turnstile> \<theta><t> is \<theta>'<t> : T"
using a1 a2 a3
proof (nominal_induct \<Gamma> t T avoiding: \<theta> \<theta>' arbitrary: \<Gamma>' rule: typing.strong_induct)
- case (T_Lam x \<Gamma> T\<^isub>1 t\<^isub>2 T\<^isub>2 \<theta> \<theta>' \<Gamma>')
+ case (T_Lam x \<Gamma> T\<^sub>1 t\<^sub>2 T\<^sub>2 \<theta> \<theta>' \<Gamma>')
have vc: "x\<sharp>\<theta>" "x\<sharp>\<theta>'" "x\<sharp>\<Gamma>" by fact+
have asm1: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>" by fact
- have ih:"\<And>\<theta> \<theta>' \<Gamma>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><t\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>2" by fact
- show "\<Gamma>' \<turnstile> \<theta><Lam [x].t\<^isub>2> is \<theta>'<Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using vc
+ have ih:"\<And>\<theta> \<theta>' \<Gamma>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^sub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><t\<^sub>2> is \<theta>'<t\<^sub>2> : T\<^sub>2" by fact
+ show "\<Gamma>' \<turnstile> \<theta><Lam [x].t\<^sub>2> is \<theta>'<Lam [x].t\<^sub>2> : T\<^sub>1\<rightarrow>T\<^sub>2" using vc
proof (simp, intro strip)
fix \<Gamma>'' s' t'
assume sub: "\<Gamma>' \<subseteq> \<Gamma>''"
- and asm2: "\<Gamma>''\<turnstile> s' is t' : T\<^isub>1"
+ and asm2: "\<Gamma>''\<turnstile> s' is t' : T\<^sub>1"
and val: "valid \<Gamma>''"
from asm1 val sub have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using logical_subst_monotonicity by blast
- with asm2 vc have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext by blast
- with ih val have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<t\<^isub>2> is ((x,t')#\<theta>')<t\<^isub>2> : T\<^isub>2" by auto
- with vc have "\<Gamma>''\<turnstile>\<theta><t\<^isub>2>[x::=s'] is \<theta>'<t\<^isub>2>[x::=t'] : T\<^isub>2" by (simp add: psubst_subst_psubst)
+ with asm2 vc have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^sub>1)#\<Gamma>" using equiv_subst_ext by blast
+ with ih val have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<t\<^sub>2> is ((x,t')#\<theta>')<t\<^sub>2> : T\<^sub>2" by auto
+ with vc have "\<Gamma>''\<turnstile>\<theta><t\<^sub>2>[x::=s'] is \<theta>'<t\<^sub>2>[x::=t'] : T\<^sub>2" by (simp add: psubst_subst_psubst)
moreover
- have "App (Lam [x].\<theta><t\<^isub>2>) s' \<leadsto> \<theta><t\<^isub>2>[x::=s']" by auto
+ have "App (Lam [x].\<theta><t\<^sub>2>) s' \<leadsto> \<theta><t\<^sub>2>[x::=s']" by auto
moreover
- have "App (Lam [x].\<theta>'<t\<^isub>2>) t' \<leadsto> \<theta>'<t\<^isub>2>[x::=t']" by auto
- ultimately show "\<Gamma>''\<turnstile> App (Lam [x].\<theta><t\<^isub>2>) s' is App (Lam [x].\<theta>'<t\<^isub>2>) t' : T\<^isub>2"
+ have "App (Lam [x].\<theta>'<t\<^sub>2>) t' \<leadsto> \<theta>'<t\<^sub>2>[x::=t']" by auto
+ ultimately show "\<Gamma>''\<turnstile> App (Lam [x].\<theta><t\<^sub>2>) s' is App (Lam [x].\<theta>'<t\<^sub>2>) t' : T\<^sub>2"
using logical_weak_head_closure by auto
qed
qed (auto)
@@ -871,68 +871,68 @@
moreover have "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T" using ih1 h v by auto
ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<u> : T" using logical_transitivity by blast
next
- case (Q_Abs x \<Gamma> T\<^isub>1 s\<^isub>2 t\<^isub>2 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
+ case (Q_Abs x \<Gamma> T\<^sub>1 s\<^sub>2 t\<^sub>2 T\<^sub>2 \<Gamma>' \<theta> \<theta>')
have fs:"x\<sharp>\<Gamma>" by fact
have fs2: "x\<sharp>\<theta>" "x\<sharp>\<theta>'" by fact+
have h2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
and h3: "valid \<Gamma>'" by fact+
- have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>2" by fact
+ have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^sub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^sub>2> is \<theta>'<t\<^sub>2> : T\<^sub>2" by fact
{
fix \<Gamma>'' s' t'
- assume "\<Gamma>' \<subseteq> \<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^isub>1" and hk: "valid \<Gamma>''"
+ assume "\<Gamma>' \<subseteq> \<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^sub>1" and hk: "valid \<Gamma>''"
then have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using h2 logical_subst_monotonicity by blast
- then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
- then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<s\<^isub>2> is ((x,t')#\<theta>')<t\<^isub>2> : T\<^isub>2" using ih hk by blast
- then have "\<Gamma>''\<turnstile> \<theta><s\<^isub>2>[x::=s'] is \<theta>'<t\<^isub>2>[x::=t'] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
- moreover have "App (Lam [x]. \<theta><s\<^isub>2>) s' \<leadsto> \<theta><s\<^isub>2>[x::=s']"
- and "App (Lam [x].\<theta>'<t\<^isub>2>) t' \<leadsto> \<theta>'<t\<^isub>2>[x::=t']" by auto
- ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<theta><s\<^isub>2>) s' is App (Lam [x].\<theta>'<t\<^isub>2>) t' : T\<^isub>2"
+ then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^sub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
+ then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<s\<^sub>2> is ((x,t')#\<theta>')<t\<^sub>2> : T\<^sub>2" using ih hk by blast
+ then have "\<Gamma>''\<turnstile> \<theta><s\<^sub>2>[x::=s'] is \<theta>'<t\<^sub>2>[x::=t'] : T\<^sub>2" using fs2 psubst_subst_psubst by auto
+ moreover have "App (Lam [x]. \<theta><s\<^sub>2>) s' \<leadsto> \<theta><s\<^sub>2>[x::=s']"
+ and "App (Lam [x].\<theta>'<t\<^sub>2>) t' \<leadsto> \<theta>'<t\<^sub>2>[x::=t']" by auto
+ ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<theta><s\<^sub>2>) s' is App (Lam [x].\<theta>'<t\<^sub>2>) t' : T\<^sub>2"
using logical_weak_head_closure by auto
}
moreover have "valid \<Gamma>'" by fact
- ultimately have "\<Gamma>' \<turnstile> Lam [x].\<theta><s\<^isub>2> is Lam [x].\<theta>'<t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
- then show "\<Gamma>' \<turnstile> \<theta><Lam [x].s\<^isub>2> is \<theta>'<Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using fs2 by auto
+ ultimately have "\<Gamma>' \<turnstile> Lam [x].\<theta><s\<^sub>2> is Lam [x].\<theta>'<t\<^sub>2> : T\<^sub>1\<rightarrow>T\<^sub>2" by auto
+ then show "\<Gamma>' \<turnstile> \<theta><Lam [x].s\<^sub>2> is \<theta>'<Lam [x].t\<^sub>2> : T\<^sub>1\<rightarrow>T\<^sub>2" using fs2 by auto
next
- case (Q_App \<Gamma> s\<^isub>1 t\<^isub>1 T\<^isub>1 T\<^isub>2 s\<^isub>2 t\<^isub>2 \<Gamma>' \<theta> \<theta>')
- then show "\<Gamma>' \<turnstile> \<theta><App s\<^isub>1 s\<^isub>2> is \<theta>'<App t\<^isub>1 t\<^isub>2> : T\<^isub>2" by auto
+ case (Q_App \<Gamma> s\<^sub>1 t\<^sub>1 T\<^sub>1 T\<^sub>2 s\<^sub>2 t\<^sub>2 \<Gamma>' \<theta> \<theta>')
+ then show "\<Gamma>' \<turnstile> \<theta><App s\<^sub>1 s\<^sub>2> is \<theta>'<App t\<^sub>1 t\<^sub>2> : T\<^sub>2" by auto
next
- case (Q_Beta x \<Gamma> s\<^isub>2 t\<^isub>2 T\<^isub>1 s12 t12 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
+ case (Q_Beta x \<Gamma> s\<^sub>2 t\<^sub>2 T\<^sub>1 s12 t12 T\<^sub>2 \<Gamma>' \<theta> \<theta>')
have h: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
and h': "valid \<Gamma>'" by fact+
have fs: "x\<sharp>\<Gamma>" by fact
have fs2: " x\<sharp>\<theta>" "x\<sharp>\<theta>'" by fact+
- have ih1: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>1" by fact
- have ih2: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s12> is \<theta>'<t12> : T\<^isub>2" by fact
- have "\<Gamma>' \<turnstile> \<theta><s\<^isub>2> is \<theta>'<t\<^isub>2> : T\<^isub>1" using ih1 h' h by auto
- then have "\<Gamma>' \<turnstile> (x,\<theta><s\<^isub>2>)#\<theta> is (x,\<theta>'<t\<^isub>2>)#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext h fs by blast
- then have "\<Gamma>' \<turnstile> ((x,\<theta><s\<^isub>2>)#\<theta>)<s12> is ((x,\<theta>'<t\<^isub>2>)#\<theta>')<t12> : T\<^isub>2" using ih2 h' by auto
- then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^isub>2>] is \<theta>'<t12>[x::=\<theta>'<t\<^isub>2>] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
- then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^isub>2>] is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 psubst_subst_propagate by auto
- moreover have "App (Lam [x].\<theta><s12>) (\<theta><s\<^isub>2>) \<leadsto> \<theta><s12>[x::=\<theta><s\<^isub>2>]" by auto
- ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<theta><s12>) (\<theta><s\<^isub>2>) is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2"
+ have ih1: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s\<^sub>2> is \<theta>'<t\<^sub>2> : T\<^sub>1" by fact
+ have ih2: "\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^sub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><s12> is \<theta>'<t12> : T\<^sub>2" by fact
+ have "\<Gamma>' \<turnstile> \<theta><s\<^sub>2> is \<theta>'<t\<^sub>2> : T\<^sub>1" using ih1 h' h by auto
+ then have "\<Gamma>' \<turnstile> (x,\<theta><s\<^sub>2>)#\<theta> is (x,\<theta>'<t\<^sub>2>)#\<theta>' over (x,T\<^sub>1)#\<Gamma>" using equiv_subst_ext h fs by blast
+ then have "\<Gamma>' \<turnstile> ((x,\<theta><s\<^sub>2>)#\<theta>)<s12> is ((x,\<theta>'<t\<^sub>2>)#\<theta>')<t12> : T\<^sub>2" using ih2 h' by auto
+ then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^sub>2>] is \<theta>'<t12>[x::=\<theta>'<t\<^sub>2>] : T\<^sub>2" using fs2 psubst_subst_psubst by auto
+ then have "\<Gamma>' \<turnstile> \<theta><s12>[x::=\<theta><s\<^sub>2>] is \<theta>'<t12[x::=t\<^sub>2]> : T\<^sub>2" using fs2 psubst_subst_propagate by auto
+ moreover have "App (Lam [x].\<theta><s12>) (\<theta><s\<^sub>2>) \<leadsto> \<theta><s12>[x::=\<theta><s\<^sub>2>]" by auto
+ ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<theta><s12>) (\<theta><s\<^sub>2>) is \<theta>'<t12[x::=t\<^sub>2]> : T\<^sub>2"
using logical_weak_head_closure' by auto
- then show "\<Gamma>' \<turnstile> \<theta><App (Lam [x].s12) s\<^isub>2> is \<theta>'<t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 by simp
+ then show "\<Gamma>' \<turnstile> \<theta><App (Lam [x].s12) s\<^sub>2> is \<theta>'<t12[x::=t\<^sub>2]> : T\<^sub>2" using fs2 by simp
next
- case (Q_Ext x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>' \<theta> \<theta>')
+ case (Q_Ext x \<Gamma> s t T\<^sub>1 T\<^sub>2 \<Gamma>' \<theta> \<theta>')
have h2: "\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma>"
and h2': "valid \<Gamma>'" by fact+
have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact+
- have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^isub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk>
- \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><App s (Var x)> is \<theta>'<App t (Var x)> : T\<^isub>2" by fact
+ have ih:"\<And>\<Gamma>' \<theta> \<theta>'. \<lbrakk>\<Gamma>' \<turnstile> \<theta> is \<theta>' over (x,T\<^sub>1)#\<Gamma>; valid \<Gamma>'\<rbrakk>
+ \<Longrightarrow> \<Gamma>' \<turnstile> \<theta><App s (Var x)> is \<theta>'<App t (Var x)> : T\<^sub>2" by fact
{
fix \<Gamma>'' s' t'
- assume hsub: "\<Gamma>' \<subseteq> \<Gamma>''" and hl: "\<Gamma>''\<turnstile> s' is t' : T\<^isub>1" and hk: "valid \<Gamma>''"
+ assume hsub: "\<Gamma>' \<subseteq> \<Gamma>''" and hl: "\<Gamma>''\<turnstile> s' is t' : T\<^sub>1" and hk: "valid \<Gamma>''"
then have "\<Gamma>'' \<turnstile> \<theta> is \<theta>' over \<Gamma>" using h2 logical_subst_monotonicity by blast
- then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
- then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<App s (Var x)> is ((x,t')#\<theta>')<App t (Var x)> : T\<^isub>2" using ih hk by blast
+ then have "\<Gamma>'' \<turnstile> (x,s')#\<theta> is (x,t')#\<theta>' over (x,T\<^sub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
+ then have "\<Gamma>'' \<turnstile> ((x,s')#\<theta>)<App s (Var x)> is ((x,t')#\<theta>')<App t (Var x)> : T\<^sub>2" using ih hk by blast
then
- have "\<Gamma>'' \<turnstile> App (((x,s')#\<theta>)<s>) (((x,s')#\<theta>)<(Var x)>) is App (((x,t')#\<theta>')<t>) (((x,t')#\<theta>')<(Var x)>) : T\<^isub>2"
+ have "\<Gamma>'' \<turnstile> App (((x,s')#\<theta>)<s>) (((x,s')#\<theta>)<(Var x)>) is App (((x,t')#\<theta>')<t>) (((x,t')#\<theta>')<(Var x)>) : T\<^sub>2"
by auto
- then have "\<Gamma>'' \<turnstile> App ((x,s')#\<theta>)<s> s' is App ((x,t')#\<theta>')<t> t' : T\<^isub>2" by auto
- then have "\<Gamma>'' \<turnstile> App (\<theta><s>) s' is App (\<theta>'<t>) t' : T\<^isub>2" using fs fresh_psubst_simp by auto
+ then have "\<Gamma>'' \<turnstile> App ((x,s')#\<theta>)<s> s' is App ((x,t')#\<theta>')<t> t' : T\<^sub>2" by auto
+ then have "\<Gamma>'' \<turnstile> App (\<theta><s>) s' is App (\<theta>'<t>) t' : T\<^sub>2" using fs fresh_psubst_simp by auto
}
moreover have "valid \<Gamma>'" by fact
- ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
+ ultimately show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : T\<^sub>1\<rightarrow>T\<^sub>2" by auto
next
case (Q_Unit \<Gamma> s t \<Gamma>' \<theta> \<theta>')
then show "\<Gamma>' \<turnstile> \<theta><s> is \<theta>'<t> : TUnit" by auto
--- a/src/HOL/Nominal/Examples/Fsub.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Fsub.thy Tue Aug 13 16:25:47 2013 +0200
@@ -54,7 +54,7 @@
abbreviation
Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("(3\<forall>_<:_./ _)" [0, 0, 10] 10)
where
- "\<forall>X<:T\<^isub>1. T\<^isub>2 \<equiv> ty.Forall X T\<^isub>2 T\<^isub>1"
+ "\<forall>X<:T\<^sub>1. T\<^sub>2 \<equiv> ty.Forall X T\<^sub>2 T\<^sub>1"
abbreviation
Abs_syn :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_:_./ _)" [0, 0, 10] 10)
@@ -402,8 +402,8 @@
where
"(Tvar X)[Y \<mapsto> T]\<^sub>\<tau> = (if X=Y then T else Tvar X)"
| "(Top)[Y \<mapsto> T]\<^sub>\<tau> = Top"
-| "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau> \<rightarrow> T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>"
-| "X\<sharp>(Y,T,T\<^isub>1) \<Longrightarrow> (\<forall>X<:T\<^isub>1. T\<^isub>2)[Y \<mapsto> T]\<^sub>\<tau> = (\<forall>X<:T\<^isub>1[Y \<mapsto> T]\<^sub>\<tau>. T\<^isub>2[Y \<mapsto> T]\<^sub>\<tau>)"
+| "(T\<^sub>1 \<rightarrow> T\<^sub>2)[Y \<mapsto> T]\<^sub>\<tau> = T\<^sub>1[Y \<mapsto> T]\<^sub>\<tau> \<rightarrow> T\<^sub>2[Y \<mapsto> T]\<^sub>\<tau>"
+| "X\<sharp>(Y,T,T\<^sub>1) \<Longrightarrow> (\<forall>X<:T\<^sub>1. T\<^sub>2)[Y \<mapsto> T]\<^sub>\<tau> = (\<forall>X<:T\<^sub>1[Y \<mapsto> T]\<^sub>\<tau>. T\<^sub>2[Y \<mapsto> T]\<^sub>\<tau>)"
apply (finite_guess)+
apply (rule TrueI)+
apply (simp add: abs_fresh)
@@ -604,8 +604,8 @@
SA_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
| SA_refl_TVar[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> ty_dom \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Tvar X <: Tvar X"
| SA_trans_TVar[intro]: "\<lbrakk>(TVarB X S) \<in> set \<Gamma>; \<Gamma> \<turnstile> S <: T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (Tvar X) <: T"
-| SA_arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^isub>1 \<rightarrow> S\<^isub>2) <: (T\<^isub>1 \<rightarrow> T\<^isub>2)"
-| SA_all[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1; ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
+| SA_arrow[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1; \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (S\<^sub>1 \<rightarrow> S\<^sub>2) <: (T\<^sub>1 \<rightarrow> T\<^sub>2)"
+| SA_all[intro]: "\<lbrakk>\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1; ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)"
lemma subtype_implies_ok:
fixes X::"tyvrs"
@@ -677,20 +677,20 @@
shows "\<Gamma> \<turnstile> T <: T"
using a b
proof(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)
- case (Forall X T\<^isub>1 T\<^isub>2)
- have ih_T\<^isub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>1 <: T\<^isub>1" by fact
- have ih_T\<^isub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^isub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" by fact
+ case (Forall X T\<^sub>1 T\<^sub>2)
+ have ih_T\<^sub>1: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^sub>1 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>1 <: T\<^sub>1" by fact
+ have ih_T\<^sub>2: "\<And>\<Gamma>. \<lbrakk>\<turnstile> \<Gamma> ok; T\<^sub>2 closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>2" by fact
have fresh_cond: "X\<sharp>\<Gamma>" by fact
hence fresh_ty_dom: "X\<sharp>(ty_dom \<Gamma>)" by (simp add: fresh_dom)
- have "(\<forall>X<:T\<^isub>2. T\<^isub>1) closed_in \<Gamma>" by fact
- hence closed\<^isub>T2: "T\<^isub>2 closed_in \<Gamma>" and closed\<^isub>T1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\<Gamma>)"
+ have "(\<forall>X<:T\<^sub>2. T\<^sub>1) closed_in \<Gamma>" by fact
+ hence closed\<^sub>T2: "T\<^sub>2 closed_in \<Gamma>" and closed\<^sub>T1: "T\<^sub>1 closed_in ((TVarB X T\<^sub>2)#\<Gamma>)"
by (auto simp add: closed_in_def ty.supp abs_supp)
have ok: "\<turnstile> \<Gamma> ok" by fact
- hence ok': "\<turnstile> ((TVarB X T\<^isub>2)#\<Gamma>) ok" using closed\<^isub>T2 fresh_ty_dom by simp
- have "\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T2 ok by simp
+ hence ok': "\<turnstile> ((TVarB X T\<^sub>2)#\<Gamma>) ok" using closed\<^sub>T2 fresh_ty_dom by simp
+ have "\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>2" using ih_T\<^sub>2 closed\<^sub>T2 ok by simp
moreover
- have "((TVarB X T\<^isub>2)#\<Gamma>) \<turnstile> T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T1 ok' by simp
- ultimately show "\<Gamma> \<turnstile> (\<forall>X<:T\<^isub>2. T\<^isub>1) <: (\<forall>X<:T\<^isub>2. T\<^isub>1)" using fresh_cond
+ have "((TVarB X T\<^sub>2)#\<Gamma>) \<turnstile> T\<^sub>1 <: T\<^sub>1" using ih_T\<^sub>1 closed\<^sub>T1 ok' by simp
+ ultimately show "\<Gamma> \<turnstile> (\<forall>X<:T\<^sub>2. T\<^sub>1) <: (\<forall>X<:T\<^sub>2. T\<^sub>1)" using fresh_cond
by (simp add: subtype_of.SA_all)
qed (auto simp add: closed_in_def ty.supp supp_atm)
@@ -772,25 +772,25 @@
hence "X \<in> ty_dom \<Delta>" using lh_drv_prem by (force dest: extends_ty_dom)
ultimately show "\<Delta> \<turnstile> Tvar X <: Tvar X" by force
next
- case (SA_arrow \<Gamma> T\<^isub>1 S\<^isub>1 S\<^isub>2 T\<^isub>2) thus "\<Delta> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by blast
+ case (SA_arrow \<Gamma> T\<^sub>1 S\<^sub>1 S\<^sub>2 T\<^sub>2) thus "\<Delta> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2" by blast
next
- case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)
+ case (SA_all \<Gamma> T\<^sub>1 S\<^sub>1 X S\<^sub>2 T\<^sub>2)
have fresh_cond: "X\<sharp>\<Delta>" by fact
hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)
- have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
- have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact
- have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
- hence closed\<^isub>T1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)
+ have ih\<^sub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact
+ have ih\<^sub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^sub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^sub>2 <: T\<^sub>2" by fact
+ have lh_drv_prem: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact
+ hence closed\<^sub>T1: "T\<^sub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)
have ok: "\<turnstile> \<Delta> ok" by fact
have ext: "\<Delta> extends \<Gamma>" by fact
- have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T1 by (simp only: extends_closed)
- hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force
+ have "T\<^sub>1 closed_in \<Delta>" using ext closed\<^sub>T1 by (simp only: extends_closed)
+ hence "\<turnstile> ((TVarB X T\<^sub>1)#\<Delta>) ok" using fresh_dom ok by force
moreover
- have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
- ultimately have "((TVarB X T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp
+ have "((TVarB X T\<^sub>1)#\<Delta>) extends ((TVarB X T\<^sub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
+ ultimately have "((TVarB X T\<^sub>1)#\<Delta>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp
moreover
- have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp
- ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all)
+ have "\<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp
+ ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)
qed
text {* In fact all ``non-binding" cases can be solved automatically: *}
@@ -802,23 +802,23 @@
shows "\<Delta> \<turnstile> S <: T"
using a b c
proof (nominal_induct \<Gamma> S T avoiding: \<Delta> rule: subtype_of.strong_induct)
- case (SA_all \<Gamma> T\<^isub>1 S\<^isub>1 X S\<^isub>2 T\<^isub>2)
+ case (SA_all \<Gamma> T\<^sub>1 S\<^sub>1 X S\<^sub>2 T\<^sub>2)
have fresh_cond: "X\<sharp>\<Delta>" by fact
hence fresh_dom: "X\<sharp>(ty_dom \<Delta>)" by (simp add: fresh_dom)
- have ih\<^isub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
- have ih\<^isub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^isub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^isub>2 <: T\<^isub>2" by fact
- have lh_drv_prem: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" by fact
- hence closed\<^isub>T1: "T\<^isub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)
+ have ih\<^sub>1: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends \<Gamma> \<Longrightarrow> \<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact
+ have ih\<^sub>2: "\<And>\<Delta>. \<turnstile> \<Delta> ok \<Longrightarrow> \<Delta> extends ((TVarB X T\<^sub>1)#\<Gamma>) \<Longrightarrow> \<Delta> \<turnstile> S\<^sub>2 <: T\<^sub>2" by fact
+ have lh_drv_prem: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" by fact
+ hence closed\<^sub>T1: "T\<^sub>1 closed_in \<Gamma>" by (simp add: subtype_implies_closed)
have ok: "\<turnstile> \<Delta> ok" by fact
have ext: "\<Delta> extends \<Gamma>" by fact
- have "T\<^isub>1 closed_in \<Delta>" using ext closed\<^isub>T1 by (simp only: extends_closed)
- hence "\<turnstile> ((TVarB X T\<^isub>1)#\<Delta>) ok" using fresh_dom ok by force
+ have "T\<^sub>1 closed_in \<Delta>" using ext closed\<^sub>T1 by (simp only: extends_closed)
+ hence "\<turnstile> ((TVarB X T\<^sub>1)#\<Delta>) ok" using fresh_dom ok by force
moreover
- have "((TVarB X T\<^isub>1)#\<Delta>) extends ((TVarB X T\<^isub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
- ultimately have "((TVarB X T\<^isub>1)#\<Delta>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using ih\<^isub>2 by simp
+ have "((TVarB X T\<^sub>1)#\<Delta>) extends ((TVarB X T\<^sub>1)#\<Gamma>)" using ext by (force simp add: extends_def)
+ ultimately have "((TVarB X T\<^sub>1)#\<Delta>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp
moreover
- have "\<Delta> \<turnstile> T\<^isub>1 <: S\<^isub>1" using ok ext ih\<^isub>1 by simp
- ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)" using ok by (force intro: SA_all)
+ have "\<Delta> \<turnstile> T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp
+ ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)
qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+
section {* Transitivity and Narrowing *}
@@ -828,13 +828,13 @@
declare ty.inject [simp add]
inductive_cases S_TopE: "\<Gamma> \<turnstile> Top <: T"
-inductive_cases S_ArrowE_left: "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T"
+inductive_cases S_ArrowE_left: "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T"
declare ty.inject [simp del]
lemma S_ForallE_left:
- shows "\<lbrakk>\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^isub>1; X\<sharp>T\<rbrakk>
- \<Longrightarrow> T = Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T = (\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2)"
+ shows "\<lbrakk>\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T; X\<sharp>\<Gamma>; X\<sharp>S\<^sub>1; X\<sharp>T\<rbrakk>
+ \<Longrightarrow> T = Top \<or> (\<exists>T\<^sub>1 T\<^sub>2. T = (\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2)"
apply(erule subtype_of.strong_cases[where X="X"])
apply(auto simp add: abs_fresh ty.inject alpha)
done
@@ -897,74 +897,74 @@
case (SA_refl_TVar \<Gamma> X)
then show "\<Gamma> \<turnstile> Tvar X <: T" by simp
next
- case (SA_arrow \<Gamma> Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2)
- then have rh_drv: "\<Gamma> \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: T" by simp
- from `Q\<^isub>1 \<rightarrow> Q\<^isub>2 = Q`
- have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto
- have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact
- have lh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact
- from rh_drv have "T=Top \<or> (\<exists>T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> \<Gamma>\<turnstile>Q\<^isub>2<:T\<^isub>2)"
+ case (SA_arrow \<Gamma> Q\<^sub>1 S\<^sub>1 S\<^sub>2 Q\<^sub>2)
+ then have rh_drv: "\<Gamma> \<turnstile> Q\<^sub>1 \<rightarrow> Q\<^sub>2 <: T" by simp
+ from `Q\<^sub>1 \<rightarrow> Q\<^sub>2 = Q`
+ have Q\<^sub>12_less: "size_ty Q\<^sub>1 < size_ty Q" "size_ty Q\<^sub>2 < size_ty Q" by auto
+ have lh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> Q\<^sub>1 <: S\<^sub>1" by fact
+ have lh_drv_prm\<^sub>2: "\<Gamma> \<turnstile> S\<^sub>2 <: Q\<^sub>2" by fact
+ from rh_drv have "T=Top \<or> (\<exists>T\<^sub>1 T\<^sub>2. T=T\<^sub>1\<rightarrow>T\<^sub>2 \<and> \<Gamma>\<turnstile>T\<^sub>1<:Q\<^sub>1 \<and> \<Gamma>\<turnstile>Q\<^sub>2<:T\<^sub>2)"
by (auto elim: S_ArrowE_left)
moreover
- have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in \<Gamma>"
- using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
- hence "(S\<^isub>1 \<rightarrow> S\<^isub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)
+ have "S\<^sub>1 closed_in \<Gamma>" and "S\<^sub>2 closed_in \<Gamma>"
+ using lh_drv_prm\<^sub>1 lh_drv_prm\<^sub>2 by (simp_all add: subtype_implies_closed)
+ hence "(S\<^sub>1 \<rightarrow> S\<^sub>2) closed_in \<Gamma>" by (simp add: closed_in_def ty.supp)
moreover
have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
moreover
- { assume "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1\<rightarrow>T\<^isub>2 \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> \<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2"
- then obtain T\<^isub>1 T\<^isub>2
- where T_inst: "T = T\<^isub>1 \<rightarrow> T\<^isub>2"
- and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"
- and rh_drv_prm\<^isub>2: "\<Gamma> \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
- from IH_trans[of "Q\<^isub>1"]
- have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using Q\<^isub>12_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp
+ { assume "\<exists>T\<^sub>1 T\<^sub>2. T = T\<^sub>1\<rightarrow>T\<^sub>2 \<and> \<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1 \<and> \<Gamma> \<turnstile> Q\<^sub>2 <: T\<^sub>2"
+ then obtain T\<^sub>1 T\<^sub>2
+ where T_inst: "T = T\<^sub>1 \<rightarrow> T\<^sub>2"
+ and rh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1"
+ and rh_drv_prm\<^sub>2: "\<Gamma> \<turnstile> Q\<^sub>2 <: T\<^sub>2" by force
+ from IH_trans[of "Q\<^sub>1"]
+ have "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" using Q\<^sub>12_less rh_drv_prm\<^sub>1 lh_drv_prm\<^sub>1 by simp
moreover
- from IH_trans[of "Q\<^isub>2"]
- have "\<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp
- ultimately have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
- then have "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" using T_inst by simp
+ from IH_trans[of "Q\<^sub>2"]
+ have "\<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2" using Q\<^sub>12_less rh_drv_prm\<^sub>2 lh_drv_prm\<^sub>2 by simp
+ ultimately have "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2" by auto
+ then have "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T" using T_inst by simp
}
- ultimately show "\<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T" by blast
+ ultimately show "\<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T" by blast
next
- case (SA_all \<Gamma> Q\<^isub>1 S\<^isub>1 X S\<^isub>2 Q\<^isub>2)
- then have rh_drv: "\<Gamma> \<turnstile> (\<forall>X<:Q\<^isub>1. Q\<^isub>2) <: T" by simp
- have lh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> Q\<^isub>1 <: S\<^isub>1" by fact
- have lh_drv_prm\<^isub>2: "((TVarB X Q\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" by fact
+ case (SA_all \<Gamma> Q\<^sub>1 S\<^sub>1 X S\<^sub>2 Q\<^sub>2)
+ then have rh_drv: "\<Gamma> \<turnstile> (\<forall>X<:Q\<^sub>1. Q\<^sub>2) <: T" by simp
+ have lh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> Q\<^sub>1 <: S\<^sub>1" by fact
+ have lh_drv_prm\<^sub>2: "((TVarB X Q\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: Q\<^sub>2" by fact
then have "X\<sharp>\<Gamma>" by (force dest: subtype_implies_ok simp add: valid_ty_dom_fresh)
- then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^isub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^isub>1
+ then have fresh_cond: "X\<sharp>\<Gamma>" "X\<sharp>Q\<^sub>1" "X\<sharp>T" using rh_drv lh_drv_prm\<^sub>1
by (simp_all add: subtype_implies_fresh)
from rh_drv
have "T = Top \<or>
- (\<exists>T\<^isub>1 T\<^isub>2. T = (\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2)"
+ (\<exists>T\<^sub>1 T\<^sub>2. T = (\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> Q\<^sub>2 <: T\<^sub>2)"
using fresh_cond by (simp add: S_ForallE_left)
moreover
- have "S\<^isub>1 closed_in \<Gamma>" and "S\<^isub>2 closed_in ((TVarB X Q\<^isub>1)#\<Gamma>)"
- using lh_drv_prm\<^isub>1 lh_drv_prm\<^isub>2 by (simp_all add: subtype_implies_closed)
- then have "(\<forall>X<:S\<^isub>1. S\<^isub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp)
+ have "S\<^sub>1 closed_in \<Gamma>" and "S\<^sub>2 closed_in ((TVarB X Q\<^sub>1)#\<Gamma>)"
+ using lh_drv_prm\<^sub>1 lh_drv_prm\<^sub>2 by (simp_all add: subtype_implies_closed)
+ then have "(\<forall>X<:S\<^sub>1. S\<^sub>2) closed_in \<Gamma>" by (force simp add: closed_in_def ty.supp abs_supp)
moreover
have "\<turnstile> \<Gamma> ok" using rh_drv by (rule subtype_implies_ok)
moreover
- { assume "\<exists>T\<^isub>1 T\<^isub>2. T=(\<forall>X<:T\<^isub>1. T\<^isub>2) \<and> \<Gamma>\<turnstile>T\<^isub>1<:Q\<^isub>1 \<and> ((TVarB X T\<^isub>1)#\<Gamma>)\<turnstile>Q\<^isub>2<:T\<^isub>2"
- then obtain T\<^isub>1 T\<^isub>2
- where T_inst: "T = (\<forall>X<:T\<^isub>1. T\<^isub>2)"
- and rh_drv_prm\<^isub>1: "\<Gamma> \<turnstile> T\<^isub>1 <: Q\<^isub>1"
- and rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> Q\<^isub>2 <: T\<^isub>2" by force
- have "(\<forall>X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact
- then have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q"
+ { assume "\<exists>T\<^sub>1 T\<^sub>2. T=(\<forall>X<:T\<^sub>1. T\<^sub>2) \<and> \<Gamma>\<turnstile>T\<^sub>1<:Q\<^sub>1 \<and> ((TVarB X T\<^sub>1)#\<Gamma>)\<turnstile>Q\<^sub>2<:T\<^sub>2"
+ then obtain T\<^sub>1 T\<^sub>2
+ where T_inst: "T = (\<forall>X<:T\<^sub>1. T\<^sub>2)"
+ and rh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> T\<^sub>1 <: Q\<^sub>1"
+ and rh_drv_prm\<^sub>2:"((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> Q\<^sub>2 <: T\<^sub>2" by force
+ have "(\<forall>X<:Q\<^sub>1. Q\<^sub>2) = Q" by fact
+ then have Q\<^sub>12_less: "size_ty Q\<^sub>1 < size_ty Q" "size_ty Q\<^sub>2 < size_ty Q"
using fresh_cond by auto
- from IH_trans[of "Q\<^isub>1"]
- have "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>12_less by blast
+ from IH_trans[of "Q\<^sub>1"]
+ have "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1" using lh_drv_prm\<^sub>1 rh_drv_prm\<^sub>1 Q\<^sub>12_less by blast
moreover
- from IH_narrow[of "Q\<^isub>1" "[]"]
- have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>12_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp
- with IH_trans[of "Q\<^isub>2"]
- have "((TVarB X T\<^isub>1)#\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 by simp
- ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: (\<forall>X<:T\<^isub>1. T\<^isub>2)"
+ from IH_narrow[of "Q\<^sub>1" "[]"]
+ have "((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: Q\<^sub>2" using Q\<^sub>12_less lh_drv_prm\<^sub>2 rh_drv_prm\<^sub>1 by simp
+ with IH_trans[of "Q\<^sub>2"]
+ have "((TVarB X T\<^sub>1)#\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2" using Q\<^sub>12_less rh_drv_prm\<^sub>2 by simp
+ ultimately have "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)"
using fresh_cond by (simp add: subtype_of.SA_all)
- hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" using T_inst by simp
+ hence "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T" using T_inst by simp
}
- ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^isub>1. S\<^isub>2) <: T" by blast
+ ultimately show "\<Gamma> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: T" by blast
qed
} note transitivity_lemma = this
@@ -992,8 +992,8 @@
then have IH_inner: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: N"
and lh_drv_prm: "(TVarB Y S) \<in> set (\<Delta>@[(TVarB X Q)]@\<Gamma>)"
and rh_drv: "\<Gamma> \<turnstile> P<:Q"
- and ok\<^isub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)
- then have ok\<^isub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok)
+ and ok\<^sub>Q: "\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok" by (simp_all add: subtype_implies_ok)
+ then have ok\<^sub>P: "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok" by (simp add: subtype_implies_ok)
show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N"
proof (cases "X=Y")
case False
@@ -1002,16 +1002,16 @@
with IH_inner show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar)
next
case True
- have memb\<^isub>XQ: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp
- have memb\<^isub>XP: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp
+ have memb\<^sub>XQ: "(TVarB X Q)\<in>set (\<Delta>@[(TVarB X Q)]@\<Gamma>)" by simp
+ have memb\<^sub>XP: "(TVarB X P)\<in>set (\<Delta>@[(TVarB X P)]@\<Gamma>)" by simp
have eq: "X=Y" by fact
- hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>XQ by (simp only: uniqueness_of_ctxt)
+ hence "S=Q" using ok\<^sub>Q lh_drv_prm memb\<^sub>XQ by (simp only: uniqueness_of_ctxt)
hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q <: N" using IH_inner by simp
moreover
have "(\<Delta>@[(TVarB X P)]@\<Gamma>) extends \<Gamma>" by (simp add: extends_def)
- hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening)
+ hence "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: Q" using rh_drv ok\<^sub>P by (simp only: weakening)
ultimately have "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> P <: N" by (simp add: transitivity_lemma)
- then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^isub>XP eq by auto
+ then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: N" using memb\<^sub>XP eq by auto
qed
next
case (SA_refl_TVar Y \<Gamma> X \<Delta>)
@@ -1024,14 +1024,14 @@
by (simp add: doms_append)
ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar)
next
- case (SA_arrow S\<^isub>1 Q\<^isub>1 Q\<^isub>2 S\<^isub>2 \<Gamma> X \<Delta>)
- then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q\<^isub>1 \<rightarrow> Q\<^isub>2 <: S\<^isub>1 \<rightarrow> S\<^isub>2" by blast
+ case (SA_arrow S\<^sub>1 Q\<^sub>1 Q\<^sub>2 S\<^sub>2 \<Gamma> X \<Delta>)
+ then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Q\<^sub>1 \<rightarrow> Q\<^sub>2 <: S\<^sub>1 \<rightarrow> S\<^sub>2" by blast
next
- case (SA_all T\<^isub>1 S\<^isub>1 Y S\<^isub>2 T\<^isub>2 \<Gamma> X \<Delta>)
- have IH_inner\<^isub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^isub>1 <: S\<^isub>1"
- and IH_inner\<^isub>2: "(((TVarB Y T\<^isub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^isub>2 <: T\<^isub>2"
+ case (SA_all T\<^sub>1 S\<^sub>1 Y S\<^sub>2 T\<^sub>2 \<Gamma> X \<Delta>)
+ have IH_inner\<^sub>1: "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> T\<^sub>1 <: S\<^sub>1"
+ and IH_inner\<^sub>2: "(((TVarB Y T\<^sub>1)#\<Delta>)@[(TVarB X P)]@\<Gamma>) \<turnstile> S\<^sub>2 <: T\<^sub>2"
by (fastforce intro: SA_all)+
- then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^isub>1. S\<^isub>2) <: (\<forall>Y<:T\<^isub>1. T\<^isub>2)" by auto
+ then show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> (\<forall>Y<:S\<^sub>1. S\<^sub>2) <: (\<forall>Y<:T\<^sub>1. T\<^sub>2)" by auto
qed
}
qed
@@ -1042,26 +1042,26 @@
typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
where
T_Var[intro]: "\<lbrakk> VarB x T \<in> set \<Gamma>; \<turnstile> \<Gamma> ok \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot> t\<^isub>2 : T\<^isub>2"
-| T_Abs[intro]: "\<lbrakk> VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1 \<rightarrow> T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<cdot> t\<^sub>2 : T\<^sub>2"
+| T_Abs[intro]: "\<lbrakk> VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>x:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \<rightarrow> T\<^sub>2"
| T_Sub[intro]: "\<lbrakk> \<Gamma> \<turnstile> t : S; \<Gamma> \<turnstile> S <: T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t : T"
-| T_TAbs[intro]:"\<lbrakk> TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>1. t\<^isub>2) : (\<forall>X<:T\<^isub>1. T\<^isub>2)"
-| T_TApp[intro]:"\<lbrakk>X\<sharp>(\<Gamma>,t\<^isub>1,T\<^isub>2); \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T\<^isub>11. T\<^isub>12); \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 : (T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>)"
+| T_TAbs[intro]:"\<lbrakk> TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>X<:T\<^sub>1. t\<^sub>2) : (\<forall>X<:T\<^sub>1. T\<^sub>2)"
+| T_TApp[intro]:"\<lbrakk>X\<sharp>(\<Gamma>,t\<^sub>1,T\<^sub>2); \<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T\<^sub>11. T\<^sub>12); \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 : (T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>)"
equivariance typing
lemma better_T_TApp:
- assumes H1: "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)"
+ assumes H1: "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T11. T12)"
and H2: "\<Gamma> \<turnstile> T2 <: T11"
- shows "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (T12[X \<mapsto> T2]\<^sub>\<tau>)"
+ shows "\<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T2 : (T12[X \<mapsto> T2]\<^sub>\<tau>)"
proof -
- obtain Y::tyvrs where Y: "Y \<sharp> (X, T12, \<Gamma>, t\<^isub>1, T2)"
+ obtain Y::tyvrs where Y: "Y \<sharp> (X, T12, \<Gamma>, t\<^sub>1, T2)"
by (rule exists_fresh) (rule fin_supp)
- then have "Y \<sharp> (\<Gamma>, t\<^isub>1, T2)" by simp
+ then have "Y \<sharp> (\<Gamma>, t\<^sub>1, T2)" by simp
moreover from Y have "(\<forall>X<:T11. T12) = (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)"
by (auto simp add: ty.inject alpha' fresh_prod fresh_atm)
- with H1 have "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" by simp
- ultimately have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot>\<^sub>\<tau> T2 : (([(Y, X)] \<bullet> T12)[Y \<mapsto> T2]\<^sub>\<tau>)" using H2
+ with H1 have "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall>Y<:T11. [(Y, X)] \<bullet> T12)" by simp
+ ultimately have "\<Gamma> \<turnstile> t\<^sub>1 \<cdot>\<^sub>\<tau> T2 : (([(Y, X)] \<bullet> T12)[Y \<mapsto> T2]\<^sub>\<tau>)" using H2
by (rule T_TApp)
with Y show ?thesis by (simp add: type_subst_rename)
qed
@@ -1118,23 +1118,23 @@
from `\<turnstile> \<Gamma> ok` and `VarB x T \<in> set \<Gamma>`
show ?case by (rule ok_imp_VarB_closed_in)
next
- case (T_App \<Gamma> t\<^isub>1 T\<^isub>1 T\<^isub>2 t\<^isub>2)
+ case (T_App \<Gamma> t\<^sub>1 T\<^sub>1 T\<^sub>2 t\<^sub>2)
then show ?case by (auto simp add: ty.supp closed_in_def)
next
- case (T_Abs x T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
- from `VarB x T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
- have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
+ case (T_Abs x T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
+ from `VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
with T_Abs show ?case by (auto simp add: ty.supp closed_in_def)
next
case (T_Sub \<Gamma> t S T)
from `\<Gamma> \<turnstile> S <: T` show ?case by (simp add: subtype_implies_closed)
next
- case (T_TAbs X T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
- from `TVarB X T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
- have "T\<^isub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
+ case (T_TAbs X T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
+ from `TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
with T_TAbs show ?case by (auto simp add: ty.supp closed_in_def abs_supp)
next
- case (T_TApp X \<Gamma> t\<^isub>1 T2 T11 T12)
+ case (T_TApp X \<Gamma> t\<^sub>1 T2 T11 T12)
then have "T12 closed_in (TVarB X T11 # \<Gamma>)"
by (auto simp add: closed_in_def ty.supp abs_supp)
moreover from T_TApp have "T2 closed_in \<Gamma>"
@@ -1161,10 +1161,10 @@
inductive
eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
where
- E_Abs : "\<lbrakk> x \<sharp> v\<^isub>2; val v\<^isub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^isub>11. t\<^isub>12) \<cdot> v\<^isub>2 \<longmapsto> t\<^isub>12[x \<mapsto> v\<^isub>2]"
+ E_Abs : "\<lbrakk> x \<sharp> v\<^sub>2; val v\<^sub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^sub>11. t\<^sub>12) \<cdot> v\<^sub>2 \<longmapsto> t\<^sub>12[x \<mapsto> v\<^sub>2]"
| E_App1 [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u"
| E_App2 [intro]: "\<lbrakk> val v; t \<longmapsto> t' \<rbrakk> \<Longrightarrow> v \<cdot> t \<longmapsto> v \<cdot> t'"
-| E_TAbs : "X \<sharp> (T\<^isub>11, T\<^isub>2) \<Longrightarrow> (\<lambda>X<:T\<^isub>11. t\<^isub>12) \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t\<^isub>12[X \<mapsto>\<^sub>\<tau> T\<^isub>2]"
+| E_TAbs : "X \<sharp> (T\<^sub>11, T\<^sub>2) \<Longrightarrow> (\<lambda>X<:T\<^sub>11. t\<^sub>12) \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2]"
| E_TApp [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot>\<^sub>\<tau> T \<longmapsto> t' \<cdot>\<^sub>\<tau> T"
lemma better_E_Abs[intro]:
@@ -1312,14 +1312,14 @@
case (T_Var x T)
then show ?case by auto
next
- case (T_App X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12)
+ case (T_App X t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12)
then show ?case by force
next
- case (T_Abs y T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>)
- then have "VarB y T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
- then have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)"
+ case (T_Abs y T\<^sub>1 t\<^sub>2 T\<^sub>2 \<Delta> \<Gamma>)
+ then have "VarB y T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp
+ then have closed: "T\<^sub>1 closed_in (\<Delta> @ \<Gamma>)"
by (auto dest: typing_ok)
- have "\<turnstile> (VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok"
+ have "\<turnstile> (VarB y T\<^sub>1 # \<Delta> @ B # \<Gamma>) ok"
apply (rule valid_cons)
apply (rule T_Abs)
apply (simp add: doms_append
@@ -1329,10 +1329,10 @@
apply (rule closed_in_weaken)
apply (rule closed)
done
- then have "\<turnstile> ((VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
- with _ have "(VarB y T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
+ then have "\<turnstile> ((VarB y T\<^sub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
+ with _ have "(VarB y T\<^sub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
by (rule T_Abs) simp
- then have "VarB y T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
+ then have "VarB y T\<^sub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp
then show ?case by (rule typing.T_Abs)
next
case (T_Sub t S T \<Delta> \<Gamma>)
@@ -1343,11 +1343,11 @@
by (rule weakening) (simp add: extends_def T_Sub)
ultimately show ?case by (rule typing.T_Sub)
next
- case (T_TAbs X T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Delta> \<Gamma>)
- from `TVarB X T\<^isub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2`
- have closed: "T\<^isub>1 closed_in (\<Delta> @ \<Gamma>)"
+ case (T_TAbs X T\<^sub>1 t\<^sub>2 T\<^sub>2 \<Delta> \<Gamma>)
+ from `TVarB X T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ have closed: "T\<^sub>1 closed_in (\<Delta> @ \<Gamma>)"
by (auto dest: typing_ok)
- have "\<turnstile> (TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma>) ok"
+ have "\<turnstile> (TVarB X T\<^sub>1 # \<Delta> @ B # \<Gamma>) ok"
apply (rule valid_consT)
apply (rule T_TAbs)
apply (simp add: doms_append
@@ -1357,14 +1357,14 @@
apply (rule closed_in_weaken)
apply (rule closed)
done
- then have "\<turnstile> ((TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
- with _ have "(TVarB X T\<^isub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
+ then have "\<turnstile> ((TVarB X T\<^sub>1 # \<Delta>) @ B # \<Gamma>) ok" by simp
+ with _ have "(TVarB X T\<^sub>1 # \<Delta>) @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
by (rule T_TAbs) simp
- then have "TVarB X T\<^isub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" by simp
+ then have "TVarB X T\<^sub>1 # \<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2" by simp
then show ?case by (rule typing.T_TAbs)
next
- case (T_TApp X t\<^isub>1 T2 T11 T12 \<Delta> \<Gamma>)
- have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^isub>1 : (\<forall>X<:T11. T12)"
+ case (T_TApp X t\<^sub>1 T2 T11 T12 \<Delta> \<Gamma>)
+ have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T11. T12)"
by (rule T_TApp refl)+
moreover from `(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
have "(\<Delta> @ B # \<Gamma>)\<turnstile>T2<:T11"
@@ -1656,17 +1656,17 @@
and "\<Gamma> \<turnstile> S' <: U'"
using H H' H''
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>x:S. s" T avoiding: x arbitrary: U U' S s rule: typing.strong_induct)
- case (T_Abs y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
- from `\<Gamma> \<turnstile> T\<^isub>1 \<rightarrow> T\<^isub>2 <: U \<rightarrow> U'`
- obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^isub>2 <: U'" using T_Abs
+ case (T_Abs y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
+ from `\<Gamma> \<turnstile> T\<^sub>1 \<rightarrow> T\<^sub>2 <: U \<rightarrow> U'`
+ obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^sub>2 <: U'" using T_Abs
by cases (simp_all add: ty.inject trm.inject alpha fresh_atm)
- from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^isub>2"
+ from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^sub>2"
by (simp add: trm.inject alpha fresh_atm)
- then have "[(y, x)] \<bullet> (VarB y S # \<Gamma>) \<turnstile> [(y, x)] \<bullet> [(y, x)] \<bullet> s : [(y, x)] \<bullet> T\<^isub>2"
+ then have "[(y, x)] \<bullet> (VarB y S # \<Gamma>) \<turnstile> [(y, x)] \<bullet> [(y, x)] \<bullet> s : [(y, x)] \<bullet> T\<^sub>2"
by (rule typing.eqvt)
moreover from T_Abs have "y \<sharp> \<Gamma>"
by (auto dest!: typing_ok simp add: fresh_trm_dom)
- ultimately have "VarB x S # \<Gamma> \<turnstile> s : T\<^isub>2" using T_Abs
+ ultimately have "VarB x S # \<Gamma> \<turnstile> s : T\<^sub>2" using T_Abs
by (perm_simp add: ty_vrs_prm_simp)
with ty1 show ?case using ty2 by (rule T_Abs)
next
@@ -1699,27 +1699,27 @@
and "(TVarB X U # \<Gamma>) \<turnstile> S' <: U'"
using H H' fresh
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct)
- case (T_TAbs Y T\<^isub>1 \<Gamma> t\<^isub>2 T\<^isub>2)
- from `TVarB Y T\<^isub>1 # \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2` have Y: "Y \<sharp> \<Gamma>"
+ case (T_TAbs Y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
+ from `TVarB Y T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2` have Y: "Y \<sharp> \<Gamma>"
by (auto dest!: typing_ok simp add: valid_ty_dom_fresh)
from `Y \<sharp> U'` and `Y \<sharp> X`
have "(\<forall>X<:U. U') = (\<forall>Y<:U. [(Y, X)] \<bullet> U')"
by (simp add: ty.inject alpha' fresh_atm)
- with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^isub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject)
- then obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "(TVarB Y U # \<Gamma>) \<turnstile> T\<^isub>2 <: ([(Y, X)] \<bullet> U')" using T_TAbs Y
+ with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^sub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject)
+ then obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "(TVarB Y U # \<Gamma>) \<turnstile> T\<^sub>2 <: ([(Y, X)] \<bullet> U')" using T_TAbs Y
by (cases rule: subtype_of.strong_cases [where X=Y]) (simp_all add: ty.inject alpha abs_fresh)
note ty1
- moreover from T_TAbs have "TVarB Y S # \<Gamma> \<turnstile> ([(Y, X)] \<bullet> s) : T\<^isub>2"
+ moreover from T_TAbs have "TVarB Y S # \<Gamma> \<turnstile> ([(Y, X)] \<bullet> s) : T\<^sub>2"
by (simp add: trm.inject alpha fresh_atm)
- then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^isub>2"
+ then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^sub>2"
by (rule typing.eqvt)
- with `X \<sharp> \<Gamma>` `X \<sharp> S` Y `Y \<sharp> S` have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2"
+ with `X \<sharp> \<Gamma>` `X \<sharp> S` Y `Y \<sharp> S` have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2"
by perm_simp
- then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^isub>2" using ty1
+ then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2" using ty1
by (rule narrow_type [of "[]", simplified])
- moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')"
+ moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')"
by (rule subtype_of.eqvt)
- with `X \<sharp> \<Gamma>` `X \<sharp> U` Y `Y \<sharp> U` have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^isub>2) <: U'"
+ with `X \<sharp> \<Gamma>` `X \<sharp> U` Y `Y \<sharp> U` have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: U'"
by perm_simp
ultimately show ?case by (rule T_TAbs)
next
@@ -1741,68 +1741,68 @@
assumes H: "\<Gamma> \<turnstile> t : T"
shows "t \<longmapsto> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : T" using H
proof (nominal_induct avoiding: t' rule: typing.strong_induct)
- case (T_App \<Gamma> t\<^isub>1 T\<^isub>11 T\<^isub>12 t\<^isub>2 t')
- obtain x::vrs where x_fresh: "x \<sharp> (\<Gamma>, t\<^isub>1 \<cdot> t\<^isub>2, t')"
+ case (T_App \<Gamma> t\<^sub>1 T\<^sub>11 T\<^sub>12 t\<^sub>2 t')
+ obtain x::vrs where x_fresh: "x \<sharp> (\<Gamma>, t\<^sub>1 \<cdot> t\<^sub>2, t')"
by (rule exists_fresh) (rule fin_supp)
- obtain X::tyvrs where "X \<sharp> (t\<^isub>1 \<cdot> t\<^isub>2, t')"
+ obtain X::tyvrs where "X \<sharp> (t\<^sub>1 \<cdot> t\<^sub>2, t')"
by (rule exists_fresh) (rule fin_supp)
- with `t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t'` show ?case
+ with `t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t'` show ?case
proof (cases rule: eval.strong_cases [where x=x and X=X])
- case (E_Abs v\<^isub>2 T\<^isub>11' t\<^isub>12)
- with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^isub>11'. t\<^isub>12) : T\<^isub>11 \<rightarrow> T\<^isub>12"
+ case (E_Abs v\<^sub>2 T\<^sub>11' t\<^sub>12)
+ with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^sub>11'. t\<^sub>12) : T\<^sub>11 \<rightarrow> T\<^sub>12"
by (simp add: trm.inject fresh_prod)
moreover from x_fresh have "x \<sharp> \<Gamma>" by simp
ultimately obtain S'
- where T\<^isub>11: "\<Gamma> \<turnstile> T\<^isub>11 <: T\<^isub>11'"
- and t\<^isub>12: "(VarB x T\<^isub>11') # \<Gamma> \<turnstile> t\<^isub>12 : S'"
- and S': "\<Gamma> \<turnstile> S' <: T\<^isub>12"
+ where T\<^sub>11: "\<Gamma> \<turnstile> T\<^sub>11 <: T\<^sub>11'"
+ and t\<^sub>12: "(VarB x T\<^sub>11') # \<Gamma> \<turnstile> t\<^sub>12 : S'"
+ and S': "\<Gamma> \<turnstile> S' <: T\<^sub>12"
by (rule Abs_type') blast
- from `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11`
- have "\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11'" using T\<^isub>11 by (rule T_Sub)
- with t\<^isub>12 have "\<Gamma> \<turnstile> t\<^isub>12[x \<mapsto> t\<^isub>2] : S'"
+ from `\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11`
+ have "\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11'" using T\<^sub>11 by (rule T_Sub)
+ with t\<^sub>12 have "\<Gamma> \<turnstile> t\<^sub>12[x \<mapsto> t\<^sub>2] : S'"
by (rule subst_type [where \<Delta>="[]", simplified])
- hence "\<Gamma> \<turnstile> t\<^isub>12[x \<mapsto> t\<^isub>2] : T\<^isub>12" using S' by (rule T_Sub)
+ hence "\<Gamma> \<turnstile> t\<^sub>12[x \<mapsto> t\<^sub>2] : T\<^sub>12" using S' by (rule T_Sub)
with E_Abs and x_fresh show ?thesis by (simp add: trm.inject fresh_prod)
next
case (E_App1 t''' t'' u)
- hence "t\<^isub>1 \<longmapsto> t''" by (simp add:trm.inject)
- hence "\<Gamma> \<turnstile> t'' : T\<^isub>11 \<rightarrow> T\<^isub>12" by (rule T_App)
- hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^isub>2 : T\<^isub>12" using `\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11`
+ hence "t\<^sub>1 \<longmapsto> t''" by (simp add:trm.inject)
+ hence "\<Gamma> \<turnstile> t'' : T\<^sub>11 \<rightarrow> T\<^sub>12" by (rule T_App)
+ hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^sub>2 : T\<^sub>12" using `\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11`
by (rule typing.T_App)
with E_App1 show ?thesis by (simp add:trm.inject)
next
case (E_App2 v t''' t'')
- hence "t\<^isub>2 \<longmapsto> t''" by (simp add:trm.inject)
- hence "\<Gamma> \<turnstile> t'' : T\<^isub>11" by (rule T_App)
- with T_App(1) have "\<Gamma> \<turnstile> t\<^isub>1 \<cdot> t'' : T\<^isub>12"
+ hence "t\<^sub>2 \<longmapsto> t''" by (simp add:trm.inject)
+ hence "\<Gamma> \<turnstile> t'' : T\<^sub>11" by (rule T_App)
+ with T_App(1) have "\<Gamma> \<turnstile> t\<^sub>1 \<cdot> t'' : T\<^sub>12"
by (rule typing.T_App)
with E_App2 show ?thesis by (simp add:trm.inject)
qed (simp_all add: fresh_prod)
next
- case (T_TApp X \<Gamma> t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12 t')
- obtain x::vrs where "x \<sharp> (t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2, t')"
+ case (T_TApp X \<Gamma> t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12 t')
+ obtain x::vrs where "x \<sharp> (t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2, t')"
by (rule exists_fresh) (rule fin_supp)
- with `t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t'`
+ with `t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t'`
show ?case
proof (cases rule: eval.strong_cases [where X=X and x=x])
- case (E_TAbs T\<^isub>11' T\<^isub>2' t\<^isub>12)
- with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^isub>11'. t\<^isub>12) : (\<forall>X<:T\<^isub>11. T\<^isub>12)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^isub>11'"
+ case (E_TAbs T\<^sub>11' T\<^sub>2' t\<^sub>12)
+ with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^sub>11'. t\<^sub>12) : (\<forall>X<:T\<^sub>11. T\<^sub>12)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^sub>11'"
by (simp_all add: trm.inject)
- moreover from `\<Gamma>\<turnstile>T\<^isub>2<:T\<^isub>11` and `X \<sharp> \<Gamma>` have "X \<sharp> T\<^isub>11"
+ moreover from `\<Gamma>\<turnstile>T\<^sub>2<:T\<^sub>11` and `X \<sharp> \<Gamma>` have "X \<sharp> T\<^sub>11"
by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed)
ultimately obtain S'
- where "TVarB X T\<^isub>11 # \<Gamma> \<turnstile> t\<^isub>12 : S'"
- and "(TVarB X T\<^isub>11 # \<Gamma>) \<turnstile> S' <: T\<^isub>12"
+ where "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : S'"
+ and "(TVarB X T\<^sub>11 # \<Gamma>) \<turnstile> S' <: T\<^sub>12"
by (rule TAbs_type') blast
- hence "TVarB X T\<^isub>11 # \<Gamma> \<turnstile> t\<^isub>12 : T\<^isub>12" by (rule T_Sub)
- hence "\<Gamma> \<turnstile> t\<^isub>12[X \<mapsto>\<^sub>\<tau> T\<^isub>2] : T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11`
+ hence "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : T\<^sub>12" by (rule T_Sub)
+ hence "\<Gamma> \<turnstile> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2] : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11`
by (rule substT_type [where D="[]", simplified])
with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject)
next
case (E_TApp t''' t'' T)
- from E_TApp have "t\<^isub>1 \<longmapsto> t''" by (simp add: trm.inject)
- then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^isub>11. T\<^isub>12)" by (rule T_TApp)
- then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^isub>2 : T\<^isub>12[X \<mapsto> T\<^isub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11`
+ from E_TApp have "t\<^sub>1 \<longmapsto> t''" by (simp add: trm.inject)
+ then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^sub>11. T\<^sub>12)" by (rule T_TApp)
+ then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11`
by (rule better_T_TApp)
with E_TApp show ?thesis by (simp add: trm.inject)
qed (simp_all add: fresh_prod)
@@ -1815,24 +1815,24 @@
qed (auto)
lemma Fun_canonical: -- {* A.14(1) *}
- assumes ty: "[] \<turnstile> v : T\<^isub>1 \<rightarrow> T\<^isub>2"
+ assumes ty: "[] \<turnstile> v : T\<^sub>1 \<rightarrow> T\<^sub>2"
shows "val v \<Longrightarrow> \<exists>x t S. v = (\<lambda>x:S. t)" using ty
-proof (induct "[]::env" v "T\<^isub>1 \<rightarrow> T\<^isub>2" arbitrary: T\<^isub>1 T\<^isub>2)
+proof (induct "[]::env" v "T\<^sub>1 \<rightarrow> T\<^sub>2" arbitrary: T\<^sub>1 T\<^sub>2)
case (T_Sub t S)
- from `[] \<turnstile> S <: T\<^isub>1 \<rightarrow> T\<^isub>2`
- obtain S\<^isub>1 S\<^isub>2 where S: "S = S\<^isub>1 \<rightarrow> S\<^isub>2"
+ from `[] \<turnstile> S <: T\<^sub>1 \<rightarrow> T\<^sub>2`
+ obtain S\<^sub>1 S\<^sub>2 where S: "S = S\<^sub>1 \<rightarrow> S\<^sub>2"
by cases (auto simp add: T_Sub)
then show ?case using `val t` by (rule T_Sub)
qed (auto)
lemma TyAll_canonical: -- {* A.14(3) *}
fixes X::tyvrs
- assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^isub>1. T\<^isub>2)"
+ assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^sub>1. T\<^sub>2)"
shows "val v \<Longrightarrow> \<exists>X t S. v = (\<lambda>X<:S. t)" using ty
-proof (induct "[]::env" v "\<forall>X<:T\<^isub>1. T\<^isub>2" arbitrary: X T\<^isub>1 T\<^isub>2)
+proof (induct "[]::env" v "\<forall>X<:T\<^sub>1. T\<^sub>2" arbitrary: X T\<^sub>1 T\<^sub>2)
case (T_Sub t S)
- from `[] \<turnstile> S <: (\<forall>X<:T\<^isub>1. T\<^isub>2)`
- obtain X S\<^isub>1 S\<^isub>2 where S: "S = (\<forall>X<:S\<^isub>1. S\<^isub>2)"
+ from `[] \<turnstile> S <: (\<forall>X<:T\<^sub>1. T\<^sub>2)`
+ obtain X S\<^sub>1 S\<^sub>2 where S: "S = (\<forall>X<:S\<^sub>1. S\<^sub>2)"
by cases (auto simp add: T_Sub)
then show ?case using T_Sub by auto
qed (auto)
@@ -1842,43 +1842,43 @@
shows "val t \<or> (\<exists>t'. t \<longmapsto> t')"
using assms
proof (induct "[]::env" t T)
- case (T_App t\<^isub>1 T\<^isub>11 T\<^isub>12 t\<^isub>2)
- hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp
+ case (T_App t\<^sub>1 T\<^sub>11 T\<^sub>12 t\<^sub>2)
+ hence "val t\<^sub>1 \<or> (\<exists>t'. t\<^sub>1 \<longmapsto> t')" by simp
thus ?case
proof
- assume t\<^isub>1_val: "val t\<^isub>1"
- with T_App obtain x t3 S where t\<^isub>1: "t\<^isub>1 = (\<lambda>x:S. t3)"
+ assume t\<^sub>1_val: "val t\<^sub>1"
+ with T_App obtain x t3 S where t\<^sub>1: "t\<^sub>1 = (\<lambda>x:S. t3)"
by (auto dest!: Fun_canonical)
- from T_App have "val t\<^isub>2 \<or> (\<exists>t'. t\<^isub>2 \<longmapsto> t')" by simp
+ from T_App have "val t\<^sub>2 \<or> (\<exists>t'. t\<^sub>2 \<longmapsto> t')" by simp
thus ?case
proof
- assume "val t\<^isub>2"
- with t\<^isub>1 have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t3[x \<mapsto> t\<^isub>2]" by auto
+ assume "val t\<^sub>2"
+ with t\<^sub>1 have "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t3[x \<mapsto> t\<^sub>2]" by auto
thus ?case by auto
next
- assume "\<exists>t'. t\<^isub>2 \<longmapsto> t'"
- then obtain t' where "t\<^isub>2 \<longmapsto> t'" by auto
- with t\<^isub>1_val have "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t\<^isub>1 \<cdot> t'" by auto
+ assume "\<exists>t'. t\<^sub>2 \<longmapsto> t'"
+ then obtain t' where "t\<^sub>2 \<longmapsto> t'" by auto
+ with t\<^sub>1_val have "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t\<^sub>1 \<cdot> t'" by auto
thus ?case by auto
qed
next
- assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'"
- then obtain t' where "t\<^isub>1 \<longmapsto> t'" by auto
- hence "t\<^isub>1 \<cdot> t\<^isub>2 \<longmapsto> t' \<cdot> t\<^isub>2" by auto
+ assume "\<exists>t'. t\<^sub>1 \<longmapsto> t'"
+ then obtain t' where "t\<^sub>1 \<longmapsto> t'" by auto
+ hence "t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t' \<cdot> t\<^sub>2" by auto
thus ?case by auto
qed
next
- case (T_TApp X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12)
- hence "val t\<^isub>1 \<or> (\<exists>t'. t\<^isub>1 \<longmapsto> t')" by simp
+ case (T_TApp X t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12)
+ hence "val t\<^sub>1 \<or> (\<exists>t'. t\<^sub>1 \<longmapsto> t')" by simp
thus ?case
proof
- assume "val t\<^isub>1"
- with T_TApp obtain x t S where "t\<^isub>1 = (\<lambda>x<:S. t)"
+ assume "val t\<^sub>1"
+ with T_TApp obtain x t S where "t\<^sub>1 = (\<lambda>x<:S. t)"
by (auto dest!: TyAll_canonical)
- hence "t\<^isub>1 \<cdot>\<^sub>\<tau> T\<^isub>2 \<longmapsto> t[x \<mapsto>\<^sub>\<tau> T\<^isub>2]" by auto
+ hence "t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t[x \<mapsto>\<^sub>\<tau> T\<^sub>2]" by auto
thus ?case by auto
next
- assume "\<exists>t'. t\<^isub>1 \<longmapsto> t'" thus ?case by auto
+ assume "\<exists>t'. t\<^sub>1 \<longmapsto> t'" thus ?case by auto
qed
qed (auto)
--- a/src/HOL/Nominal/Examples/Lam_Funs.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Tue Aug 13 16:25:47 2013 +0200
@@ -79,7 +79,7 @@
psubst :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" ("_<_>" [95,95] 105)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
-| "\<theta><(App e\<^isub>1 e\<^isub>2)> = App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>)"
+| "\<theta><(App e\<^sub>1 e\<^sub>2)> = App (\<theta><e\<^sub>1>) (\<theta><e\<^sub>2>)"
| "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].e)> = Lam [x].(\<theta><e>)"
apply(finite_guess)+
apply(rule TrueI)+
--- a/src/HOL/Nominal/Examples/Pattern.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Pattern.thy Tue Aug 13 16:25:47 2013 +0200
@@ -144,7 +144,7 @@
ptyping :: "pat \<Rightarrow> ty \<Rightarrow> ctx \<Rightarrow> bool" ("\<turnstile> _ : _ \<Rightarrow> _" [60, 60, 60] 60)
where
PVar: "\<turnstile> PVar x T : T \<Rightarrow> [(x, T)]"
-| PTuple: "\<turnstile> p : T \<Rightarrow> \<Delta>\<^isub>1 \<Longrightarrow> \<turnstile> q : U \<Rightarrow> \<Delta>\<^isub>2 \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> : T \<otimes> U \<Rightarrow> \<Delta>\<^isub>2 @ \<Delta>\<^isub>1"
+| PTuple: "\<turnstile> p : T \<Rightarrow> \<Delta>\<^sub>1 \<Longrightarrow> \<turnstile> q : U \<Rightarrow> \<Delta>\<^sub>2 \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> : T \<otimes> U \<Rightarrow> \<Delta>\<^sub>2 @ \<Delta>\<^sub>1"
lemma pat_vars_ptyping:
assumes "\<turnstile> p : T \<Rightarrow> \<Delta>"
@@ -168,7 +168,7 @@
abbreviation
"sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _")
where
- "\<Gamma>\<^isub>1 \<sqsubseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^isub>1 \<longrightarrow> x \<in> set \<Gamma>\<^isub>2"
+ "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
abbreviation
Let_syn :: "pat \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" ("(LET (_ =/ _)/ IN (_))" 10)
@@ -213,8 +213,8 @@
by (auto simp add: fresh_star_def pat_var)
lemma valid_app_mono:
- assumes "valid (\<Delta> @ \<Gamma>\<^isub>1)" and "(supp \<Delta>::name set) \<sharp>* \<Gamma>\<^isub>2" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<sqsubseteq> \<Gamma>\<^isub>2"
- shows "valid (\<Delta> @ \<Gamma>\<^isub>2)" using assms
+ assumes "valid (\<Delta> @ \<Gamma>\<^sub>1)" and "(supp \<Delta>::name set) \<sharp>* \<Gamma>\<^sub>2" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2"
+ shows "valid (\<Delta> @ \<Gamma>\<^sub>2)" using assms
by (induct \<Delta>)
(auto simp add: supp_list_cons fresh_star_Un_elim supp_prod
fresh_list_append supp_atm fresh_star_insert_elim fresh_star_empty_elim)
@@ -253,14 +253,14 @@
qed
lemma weakening:
- assumes "\<Gamma>\<^isub>1 \<turnstile> t : T" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<sqsubseteq> \<Gamma>\<^isub>2"
- shows "\<Gamma>\<^isub>2 \<turnstile> t : T" using assms
- apply (nominal_induct \<Gamma>\<^isub>1 t T avoiding: \<Gamma>\<^isub>2 rule: typing.strong_induct)
+ assumes "\<Gamma>\<^sub>1 \<turnstile> t : T" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2"
+ shows "\<Gamma>\<^sub>2 \<turnstile> t : T" using assms
+ apply (nominal_induct \<Gamma>\<^sub>1 t T avoiding: \<Gamma>\<^sub>2 rule: typing.strong_induct)
apply auto
- apply (drule_tac x="(x, T) # \<Gamma>\<^isub>2" in meta_spec)
+ apply (drule_tac x="(x, T) # \<Gamma>\<^sub>2" in meta_spec)
apply (auto intro: valid_typing)
- apply (drule_tac x="\<Gamma>\<^isub>2" in meta_spec)
- apply (drule_tac x="\<Delta> @ \<Gamma>\<^isub>2" in meta_spec)
+ apply (drule_tac x="\<Gamma>\<^sub>2" in meta_spec)
+ apply (drule_tac x="\<Delta> @ \<Gamma>\<^sub>2" in meta_spec)
apply (auto intro: valid_typing)
apply (rule typing.Let)
apply assumption+
@@ -379,8 +379,8 @@
(simp_all add: fresh_list_nil fresh_list_cons psubst_forget)
lemma psubst_append:
- "(supp (map fst (\<theta>\<^isub>1 @ \<theta>\<^isub>2))::name set) \<sharp>* map snd (\<theta>\<^isub>1 @ \<theta>\<^isub>2) \<Longrightarrow> (\<theta>\<^isub>1 @ \<theta>\<^isub>2)\<lparr>t\<rparr> = \<theta>\<^isub>2\<lparr>\<theta>\<^isub>1\<lparr>t\<rparr>\<rparr>"
- by (induct \<theta>\<^isub>1 arbitrary: t)
+ "(supp (map fst (\<theta>\<^sub>1 @ \<theta>\<^sub>2))::name set) \<sharp>* map snd (\<theta>\<^sub>1 @ \<theta>\<^sub>2) \<Longrightarrow> (\<theta>\<^sub>1 @ \<theta>\<^sub>2)\<lparr>t\<rparr> = \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr>"
+ by (induct \<theta>\<^sub>1 arbitrary: t)
(simp_all add: psubst_nil split_paired_all supp_list_cons psubst_cons fresh_star_def
fresh_list_cons fresh_list_append supp_list_append)
@@ -424,12 +424,12 @@
then show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T" using ineq valid by auto
qed
next
- case (Tuple t\<^isub>1 T\<^isub>1 t\<^isub>2 T\<^isub>2)
+ case (Tuple t\<^sub>1 T\<^sub>1 t\<^sub>2 T\<^sub>2)
from refl `\<Gamma> \<turnstile> u : U`
- have "\<Delta> @ \<Gamma> \<turnstile> t\<^isub>1[x\<mapsto>u] : T\<^isub>1" by (rule Tuple)
+ have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T\<^sub>1" by (rule Tuple)
moreover from refl `\<Gamma> \<turnstile> u : U`
- have "\<Delta> @ \<Gamma> \<turnstile> t\<^isub>2[x\<mapsto>u] : T\<^isub>2" by (rule Tuple)
- ultimately have "\<Delta> @ \<Gamma> \<turnstile> \<langle>t\<^isub>1[x\<mapsto>u], t\<^isub>2[x\<mapsto>u]\<rangle> : T\<^isub>1 \<otimes> T\<^isub>2" ..
+ have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T\<^sub>2" by (rule Tuple)
+ ultimately have "\<Delta> @ \<Gamma> \<turnstile> \<langle>t\<^sub>1[x\<mapsto>u], t\<^sub>2[x\<mapsto>u]\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" ..
then show ?case by simp
next
case (Let p t T \<Delta>' s S)
@@ -456,12 +456,12 @@
by (simp add: fresh_list_nil fresh_list_cons)
ultimately show ?case by simp
next
- case (App t\<^isub>1 T S t\<^isub>2)
+ case (App t\<^sub>1 T S t\<^sub>2)
from refl `\<Gamma> \<turnstile> u : U`
- have "\<Delta> @ \<Gamma> \<turnstile> t\<^isub>1[x\<mapsto>u] : T \<rightarrow> S" by (rule App)
+ have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T \<rightarrow> S" by (rule App)
moreover from refl `\<Gamma> \<turnstile> u : U`
- have "\<Delta> @ \<Gamma> \<turnstile> t\<^isub>2[x\<mapsto>u] : T" by (rule App)
- ultimately have "\<Delta> @ \<Gamma> \<turnstile> (t\<^isub>1[x\<mapsto>u]) \<cdot> (t\<^isub>2[x\<mapsto>u]) : S"
+ have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T" by (rule App)
+ ultimately have "\<Delta> @ \<Gamma> \<turnstile> (t\<^sub>1[x\<mapsto>u]) \<cdot> (t\<^sub>2[x\<mapsto>u]) : S"
by (rule typing.App)
then show ?case by simp
qed
@@ -482,42 +482,42 @@
lemma match_type_aux:
assumes "\<turnstile> p : U \<Rightarrow> \<Delta>"
- and "\<Gamma>\<^isub>2 \<turnstile> u : U"
- and "\<Gamma>\<^isub>1 @ \<Delta> @ \<Gamma>\<^isub>2 \<turnstile> t : T"
+ and "\<Gamma>\<^sub>2 \<turnstile> u : U"
+ and "\<Gamma>\<^sub>1 @ \<Delta> @ \<Gamma>\<^sub>2 \<turnstile> t : T"
and "\<turnstile> p \<rhd> u \<Rightarrow> \<theta>"
and "(supp p::name set) \<sharp>* u"
- shows "\<Gamma>\<^isub>1 @ \<Gamma>\<^isub>2 \<turnstile> \<theta>\<lparr>t\<rparr> : T" using assms
-proof (induct arbitrary: \<Gamma>\<^isub>1 \<Gamma>\<^isub>2 t u T \<theta>)
+ shows "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<lparr>t\<rparr> : T" using assms
+proof (induct arbitrary: \<Gamma>\<^sub>1 \<Gamma>\<^sub>2 t u T \<theta>)
case (PVar x U)
- from `\<Gamma>\<^isub>1 @ [(x, U)] @ \<Gamma>\<^isub>2 \<turnstile> t : T` `\<Gamma>\<^isub>2 \<turnstile> u : U`
- have "\<Gamma>\<^isub>1 @ \<Gamma>\<^isub>2 \<turnstile> t[x\<mapsto>u] : T" by (rule subst_type_aux)
+ from `\<Gamma>\<^sub>1 @ [(x, U)] @ \<Gamma>\<^sub>2 \<turnstile> t : T` `\<Gamma>\<^sub>2 \<turnstile> u : U`
+ have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t[x\<mapsto>u] : T" by (rule subst_type_aux)
moreover from `\<turnstile> PVar x U \<rhd> u \<Rightarrow> \<theta>` have "\<theta> = [(x, u)]"
by cases simp_all
ultimately show ?case by simp
next
- case (PTuple p S \<Delta>\<^isub>1 q U \<Delta>\<^isub>2)
- from `\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>` obtain u\<^isub>1 u\<^isub>2 \<theta>\<^isub>1 \<theta>\<^isub>2
- where u: "u = \<langle>u\<^isub>1, u\<^isub>2\<rangle>" and \<theta>: "\<theta> = \<theta>\<^isub>1 @ \<theta>\<^isub>2"
- and p: "\<turnstile> p \<rhd> u\<^isub>1 \<Rightarrow> \<theta>\<^isub>1" and q: "\<turnstile> q \<rhd> u\<^isub>2 \<Rightarrow> \<theta>\<^isub>2"
+ case (PTuple p S \<Delta>\<^sub>1 q U \<Delta>\<^sub>2)
+ from `\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>` obtain u\<^sub>1 u\<^sub>2 \<theta>\<^sub>1 \<theta>\<^sub>2
+ where u: "u = \<langle>u\<^sub>1, u\<^sub>2\<rangle>" and \<theta>: "\<theta> = \<theta>\<^sub>1 @ \<theta>\<^sub>2"
+ and p: "\<turnstile> p \<rhd> u\<^sub>1 \<Rightarrow> \<theta>\<^sub>1" and q: "\<turnstile> q \<rhd> u\<^sub>2 \<Rightarrow> \<theta>\<^sub>2"
by cases simp_all
- with PTuple have "\<Gamma>\<^isub>2 \<turnstile> \<langle>u\<^isub>1, u\<^isub>2\<rangle> : S \<otimes> U" by simp
- then obtain u\<^isub>1: "\<Gamma>\<^isub>2 \<turnstile> u\<^isub>1 : S" and u\<^isub>2: "\<Gamma>\<^isub>2 \<turnstile> u\<^isub>2 : U"
+ with PTuple have "\<Gamma>\<^sub>2 \<turnstile> \<langle>u\<^sub>1, u\<^sub>2\<rangle> : S \<otimes> U" by simp
+ then obtain u\<^sub>1: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>1 : S" and u\<^sub>2: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>2 : U"
by cases (simp_all add: ty.inject trm.inject)
- note u\<^isub>1
- moreover from `\<Gamma>\<^isub>1 @ (\<Delta>\<^isub>2 @ \<Delta>\<^isub>1) @ \<Gamma>\<^isub>2 \<turnstile> t : T`
- have "(\<Gamma>\<^isub>1 @ \<Delta>\<^isub>2) @ \<Delta>\<^isub>1 @ \<Gamma>\<^isub>2 \<turnstile> t : T" by simp
+ note u\<^sub>1
+ moreover from `\<Gamma>\<^sub>1 @ (\<Delta>\<^sub>2 @ \<Delta>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile> t : T`
+ have "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Delta>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t : T" by simp
moreover note p
moreover from `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u` and u
- have "(supp p::name set) \<sharp>* u\<^isub>1" by (simp add: fresh_star_def)
- ultimately have \<theta>\<^isub>1: "(\<Gamma>\<^isub>1 @ \<Delta>\<^isub>2) @ \<Gamma>\<^isub>2 \<turnstile> \<theta>\<^isub>1\<lparr>t\<rparr> : T"
+ have "(supp p::name set) \<sharp>* u\<^sub>1" by (simp add: fresh_star_def)
+ ultimately have \<theta>\<^sub>1: "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T"
by (rule PTuple)
- note u\<^isub>2
- moreover from \<theta>\<^isub>1
- have "\<Gamma>\<^isub>1 @ \<Delta>\<^isub>2 @ \<Gamma>\<^isub>2 \<turnstile> \<theta>\<^isub>1\<lparr>t\<rparr> : T" by simp
+ note u\<^sub>2
+ moreover from \<theta>\<^sub>1
+ have "\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T" by simp
moreover note q
moreover from `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u` and u
- have "(supp q::name set) \<sharp>* u\<^isub>2" by (simp add: fresh_star_def)
- ultimately have "\<Gamma>\<^isub>1 @ \<Gamma>\<^isub>2 \<turnstile> \<theta>\<^isub>2\<lparr>\<theta>\<^isub>1\<lparr>t\<rparr>\<rparr> : T"
+ have "(supp q::name set) \<sharp>* u\<^sub>2" by (simp add: fresh_star_def)
+ ultimately have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr> : T"
by (rule PTuple)
moreover from `\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>` `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u`
have "(supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>"
@@ -525,7 +525,7 @@
ultimately show ?case using \<theta> by (simp add: psubst_append)
qed
-lemmas match_type = match_type_aux [where \<Gamma>\<^isub>1="[]", simplified]
+lemmas match_type = match_type_aux [where \<Gamma>\<^sub>1="[]", simplified]
inductive eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
where
@@ -680,78 +680,78 @@
then show ?thesis ..
qed
next
- case (PTuple p\<^isub>1 p\<^isub>2)
- with PVar have "ty_size (pat_type p\<^isub>1) < ty_size T" by simp
- then have "Bind T x t \<noteq> (\<lambda>[p\<^isub>1]. \<lambda>[p\<^isub>2]. u)"
+ case (PTuple p\<^sub>1 p\<^sub>2)
+ with PVar have "ty_size (pat_type p\<^sub>1) < ty_size T" by simp
+ then have "Bind T x t \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)"
by (rule bind_tuple_ineq)
moreover from PTuple PVar
- have "Bind T x t = (\<lambda>[p\<^isub>1]. \<lambda>[p\<^isub>2]. u)" by simp
+ have "Bind T x t = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. u)" by simp
ultimately show ?thesis ..
qed
next
- case (PTuple p\<^isub>1 p\<^isub>2)
+ case (PTuple p\<^sub>1 p\<^sub>2)
note PTuple' = this
show ?case
proof (cases q)
case (PVar x T)
- with PTuple have "ty_size (pat_type p\<^isub>1) < ty_size T" by auto
- then have "Bind T x u \<noteq> (\<lambda>[p\<^isub>1]. \<lambda>[p\<^isub>2]. t)"
+ with PTuple have "ty_size (pat_type p\<^sub>1) < ty_size T" by auto
+ then have "Bind T x u \<noteq> (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)"
by (rule bind_tuple_ineq)
moreover from PTuple PVar
- have "Bind T x u = (\<lambda>[p\<^isub>1]. \<lambda>[p\<^isub>2]. t)" by simp
+ have "Bind T x u = (\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t)" by simp
ultimately show ?thesis ..
next
- case (PTuple p\<^isub>1' p\<^isub>2')
- with PTuple' have "(\<lambda>[p\<^isub>1]. \<lambda>[p\<^isub>2]. t) = (\<lambda>[p\<^isub>1']. \<lambda>[p\<^isub>2']. u)" by simp
- moreover from PTuple PTuple' have "pat_type p\<^isub>1 = pat_type p\<^isub>1'"
+ case (PTuple p\<^sub>1' p\<^sub>2')
+ with PTuple' have "(\<lambda>[p\<^sub>1]. \<lambda>[p\<^sub>2]. t) = (\<lambda>[p\<^sub>1']. \<lambda>[p\<^sub>2']. u)" by simp
+ moreover from PTuple PTuple' have "pat_type p\<^sub>1 = pat_type p\<^sub>1'"
by (simp add: ty.inject)
- moreover from PTuple' have "distinct (pat_vars p\<^isub>1)" by simp
- moreover from PTuple PTuple' have "distinct (pat_vars p\<^isub>1')" by simp
- ultimately have "\<exists>pi::name prm. p\<^isub>1 = pi \<bullet> p\<^isub>1' \<and>
- (\<lambda>[p\<^isub>2]. t) = pi \<bullet> (\<lambda>[p\<^isub>2']. u) \<and>
- set pi \<subseteq> (supp p\<^isub>1 \<union> supp p\<^isub>1') \<times> (supp p\<^isub>1 \<union> supp p\<^isub>1')"
+ moreover from PTuple' have "distinct (pat_vars p\<^sub>1)" by simp
+ moreover from PTuple PTuple' have "distinct (pat_vars p\<^sub>1')" by simp
+ ultimately have "\<exists>pi::name prm. p\<^sub>1 = pi \<bullet> p\<^sub>1' \<and>
+ (\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u) \<and>
+ set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')"
by (rule PTuple')
then obtain pi::"name prm" where
- "p\<^isub>1 = pi \<bullet> p\<^isub>1'" "(\<lambda>[p\<^isub>2]. t) = pi \<bullet> (\<lambda>[p\<^isub>2']. u)" and
- pi: "set pi \<subseteq> (supp p\<^isub>1 \<union> supp p\<^isub>1') \<times> (supp p\<^isub>1 \<union> supp p\<^isub>1')" by auto
- from `(\<lambda>[p\<^isub>2]. t) = pi \<bullet> (\<lambda>[p\<^isub>2']. u)`
- have "(\<lambda>[p\<^isub>2]. t) = (\<lambda>[pi \<bullet> p\<^isub>2']. pi \<bullet> u)"
+ "p\<^sub>1 = pi \<bullet> p\<^sub>1'" "(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)" and
+ pi: "set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')" by auto
+ from `(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)`
+ have "(\<lambda>[p\<^sub>2]. t) = (\<lambda>[pi \<bullet> p\<^sub>2']. pi \<bullet> u)"
by (simp add: eqvts)
- moreover from PTuple PTuple' have "pat_type p\<^isub>2 = pat_type (pi \<bullet> p\<^isub>2')"
+ moreover from PTuple PTuple' have "pat_type p\<^sub>2 = pat_type (pi \<bullet> p\<^sub>2')"
by (simp add: ty.inject pat_type_perm_eq)
- moreover from PTuple' have "distinct (pat_vars p\<^isub>2)" by simp
- moreover from PTuple PTuple' have "distinct (pat_vars (pi \<bullet> p\<^isub>2'))"
+ moreover from PTuple' have "distinct (pat_vars p\<^sub>2)" by simp
+ moreover from PTuple PTuple' have "distinct (pat_vars (pi \<bullet> p\<^sub>2'))"
by (simp add: pat_vars_eqvt [symmetric] distinct_eqvt [symmetric])
- ultimately have "\<exists>pi'::name prm. p\<^isub>2 = pi' \<bullet> pi \<bullet> p\<^isub>2' \<and>
+ ultimately have "\<exists>pi'::name prm. p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2' \<and>
t = pi' \<bullet> pi \<bullet> u \<and>
- set pi' \<subseteq> (supp p\<^isub>2 \<union> supp (pi \<bullet> p\<^isub>2')) \<times> (supp p\<^isub>2 \<union> supp (pi \<bullet> p\<^isub>2'))"
+ set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))"
by (rule PTuple')
then obtain pi'::"name prm" where
- "p\<^isub>2 = pi' \<bullet> pi \<bullet> p\<^isub>2'" "t = pi' \<bullet> pi \<bullet> u" and
- pi': "set pi' \<subseteq> (supp p\<^isub>2 \<union> supp (pi \<bullet> p\<^isub>2')) \<times>
- (supp p\<^isub>2 \<union> supp (pi \<bullet> p\<^isub>2'))" by auto
- from PTuple PTuple' have "pi \<bullet> distinct (pat_vars \<langle>\<langle>p\<^isub>1', p\<^isub>2'\<rangle>\<rangle>)" by simp
- then have "distinct (pat_vars \<langle>\<langle>pi \<bullet> p\<^isub>1', pi \<bullet> p\<^isub>2'\<rangle>\<rangle>)" by (simp only: eqvts)
- with `p\<^isub>1 = pi \<bullet> p\<^isub>1'` PTuple'
- have fresh: "(supp p\<^isub>2 \<union> supp (pi \<bullet> p\<^isub>2') :: name set) \<sharp>* p\<^isub>1"
+ "p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'" "t = pi' \<bullet> pi \<bullet> u" and
+ pi': "set pi' \<subseteq> (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2')) \<times>
+ (supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))" by auto
+ from PTuple PTuple' have "pi \<bullet> distinct (pat_vars \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)" by simp
+ then have "distinct (pat_vars \<langle>\<langle>pi \<bullet> p\<^sub>1', pi \<bullet> p\<^sub>2'\<rangle>\<rangle>)" by (simp only: eqvts)
+ with `p\<^sub>1 = pi \<bullet> p\<^sub>1'` PTuple'
+ have fresh: "(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2') :: name set) \<sharp>* p\<^sub>1"
by (auto simp add: set_pat_vars_supp fresh_star_def fresh_def eqvts)
- from `p\<^isub>1 = pi \<bullet> p\<^isub>1'` have "pi' \<bullet> (p\<^isub>1 = pi \<bullet> p\<^isub>1')" by (rule perm_boolI)
+ from `p\<^sub>1 = pi \<bullet> p\<^sub>1'` have "pi' \<bullet> (p\<^sub>1 = pi \<bullet> p\<^sub>1')" by (rule perm_boolI)
with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' fresh fresh]
- have "p\<^isub>1 = pi' \<bullet> pi \<bullet> p\<^isub>1'" by (simp add: eqvts)
- with `p\<^isub>2 = pi' \<bullet> pi \<bullet> p\<^isub>2'` have "\<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^isub>1', p\<^isub>2'\<rangle>\<rangle>"
+ have "p\<^sub>1 = pi' \<bullet> pi \<bullet> p\<^sub>1'" by (simp add: eqvts)
+ with `p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'` have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>"
by (simp add: pt_name2)
moreover
- have "((supp p\<^isub>2 \<union> (pi \<bullet> supp p\<^isub>2')) \<times> (supp p\<^isub>2 \<union> (pi \<bullet> supp p\<^isub>2'))::(name \<times> name) set) \<subseteq>
- (supp p\<^isub>2 \<union> (supp p\<^isub>1 \<union> supp p\<^isub>1' \<union> supp p\<^isub>2')) \<times> (supp p\<^isub>2 \<union> (supp p\<^isub>1 \<union> supp p\<^isub>1' \<union> supp p\<^isub>2'))"
+ have "((supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2'))::(name \<times> name) set) \<subseteq>
+ (supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (supp p\<^sub>1 \<union> supp p\<^sub>1' \<union> supp p\<^sub>2'))"
by (rule subset_refl Sigma_mono Un_mono perm_cases [OF pi])+
with pi' have "set pi' \<subseteq> \<dots>" by (simp add: supp_eqvt [symmetric])
- with pi have "set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^isub>1', p\<^isub>2'\<rangle>\<rangle>) \<times>
- (supp \<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^isub>1', p\<^isub>2'\<rangle>\<rangle>)"
+ with pi have "set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>) \<times>
+ (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)"
by (simp add: Sigma_Un_distrib1 Sigma_Un_distrib2 Un_ac) blast
moreover note `t = pi' \<bullet> pi \<bullet> u`
- ultimately have "\<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> q \<and> t = (pi' @ pi) \<bullet> u \<and>
- set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> \<union> supp q) \<times>
- (supp \<langle>\<langle>p\<^isub>1, p\<^isub>2\<rangle>\<rangle> \<union> supp q)" using PTuple
+ ultimately have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> q \<and> t = (pi' @ pi) \<bullet> u \<and>
+ set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q) \<times>
+ (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q)" using PTuple
by (simp add: pt_name2)
then show ?thesis ..
qed
@@ -805,22 +805,22 @@
shows "\<Gamma> \<turnstile> t' : T" using assms
proof (nominal_induct avoiding: \<Gamma> T rule: eval.strong_induct)
case (TupleL t t' u)
- from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^isub>1 T\<^isub>2
- where "T = T\<^isub>1 \<otimes> T\<^isub>2" "\<Gamma> \<turnstile> t : T\<^isub>1" "\<Gamma> \<turnstile> u : T\<^isub>2"
+ from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^sub>1 T\<^sub>2
+ where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
by cases (simp_all add: trm.inject)
- from `\<Gamma> \<turnstile> t : T\<^isub>1` have "\<Gamma> \<turnstile> t' : T\<^isub>1" by (rule TupleL)
- then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^isub>1 \<otimes> T\<^isub>2" using `\<Gamma> \<turnstile> u : T\<^isub>2`
+ from `\<Gamma> \<turnstile> t : T\<^sub>1` have "\<Gamma> \<turnstile> t' : T\<^sub>1" by (rule TupleL)
+ then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" using `\<Gamma> \<turnstile> u : T\<^sub>2`
by (rule Tuple)
- with `T = T\<^isub>1 \<otimes> T\<^isub>2` show ?case by simp
+ with `T = T\<^sub>1 \<otimes> T\<^sub>2` show ?case by simp
next
case (TupleR u u' t)
- from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^isub>1 T\<^isub>2
- where "T = T\<^isub>1 \<otimes> T\<^isub>2" "\<Gamma> \<turnstile> t : T\<^isub>1" "\<Gamma> \<turnstile> u : T\<^isub>2"
+ from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^sub>1 T\<^sub>2
+ where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
by cases (simp_all add: trm.inject)
- from `\<Gamma> \<turnstile> u : T\<^isub>2` have "\<Gamma> \<turnstile> u' : T\<^isub>2" by (rule TupleR)
- with `\<Gamma> \<turnstile> t : T\<^isub>1` have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^isub>1 \<otimes> T\<^isub>2"
+ from `\<Gamma> \<turnstile> u : T\<^sub>2` have "\<Gamma> \<turnstile> u' : T\<^sub>2" by (rule TupleR)
+ with `\<Gamma> \<turnstile> t : T\<^sub>1` have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2"
by (rule Tuple)
- with `T = T\<^isub>1 \<otimes> T\<^isub>2` show ?case by simp
+ with `T = T\<^sub>1 \<otimes> T\<^sub>2` show ?case by simp
next
case (Abs t t' x S)
from `\<Gamma> \<turnstile> (\<lambda>x:S. t) : T` `x \<sharp> \<Gamma>` obtain U where
--- a/src/HOL/Nominal/Examples/SN.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/SN.thy Tue Aug 13 16:25:47 2013 +0200
@@ -71,12 +71,12 @@
(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
inductive
- Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
+ Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
where
- b1[intro!]: "s1 \<longrightarrow>\<^isub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^isub>\<beta> App s2 t"
-| b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^isub>\<beta> App t s2"
-| b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> Lam [a].s1 \<longrightarrow>\<^isub>\<beta> Lam [a].s2"
-| b4[intro!]: "a\<sharp>s2 \<Longrightarrow> App (Lam [a].s1) s2\<longrightarrow>\<^isub>\<beta> (s1[a::=s2])"
+ b1[intro!]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^sub>\<beta> App s2 t"
+| b2[intro!]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
+| b3[intro!]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> Lam [a].s1 \<longrightarrow>\<^sub>\<beta> Lam [a].s2"
+| b4[intro!]: "a\<sharp>s2 \<Longrightarrow> App (Lam [a].s1) s2\<longrightarrow>\<^sub>\<beta> (s1[a::=s2])"
equivariance Beta
@@ -85,7 +85,7 @@
lemma beta_preserves_fresh:
fixes a::"name"
- assumes a: "t\<longrightarrow>\<^isub>\<beta> s"
+ assumes a: "t\<longrightarrow>\<^sub>\<beta> s"
shows "a\<sharp>t \<Longrightarrow> a\<sharp>s"
using a
apply(nominal_induct t s avoiding: a rule: Beta.strong_induct)
@@ -93,8 +93,8 @@
done
lemma beta_abs:
- assumes a: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'"
- shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''"
+ assumes a: "Lam [a].t\<longrightarrow>\<^sub>\<beta> t'"
+ shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^sub>\<beta> t''"
proof -
have "a\<sharp>Lam [a].t" by (simp add: abs_fresh)
with a have "a\<sharp>t'" by (simp add: beta_preserves_fresh)
@@ -104,8 +104,8 @@
qed
lemma beta_subst:
- assumes a: "M \<longrightarrow>\<^isub>\<beta> M'"
- shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]"
+ assumes a: "M \<longrightarrow>\<^sub>\<beta> M'"
+ shows "M[x::=N]\<longrightarrow>\<^sub>\<beta> M'[x::=N]"
using a
by (nominal_induct M M' avoiding: x N rule: Beta.strong_induct)
(auto simp add: fresh_atm subst_lemma fresh_fact)
@@ -159,13 +159,13 @@
subsection {* a fact about beta *}
definition "NORMAL" :: "lam \<Rightarrow> bool" where
- "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')"
+ "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^sub>\<beta> t')"
lemma NORMAL_Var:
shows "NORMAL (Var a)"
proof -
- { assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
- then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast
+ { assume "\<exists>t'. (Var a) \<longrightarrow>\<^sub>\<beta> t'"
+ then obtain t' where "(Var a) \<longrightarrow>\<^sub>\<beta> t'" by blast
hence False by (cases) (auto)
}
thus "NORMAL (Var a)" by (auto simp add: NORMAL_def)
@@ -175,10 +175,10 @@
inductive
SN :: "lam \<Rightarrow> bool"
where
- SN_intro: "(\<And>t'. t \<longrightarrow>\<^isub>\<beta> t' \<Longrightarrow> SN t') \<Longrightarrow> SN t"
+ SN_intro: "(\<And>t'. t \<longrightarrow>\<^sub>\<beta> t' \<Longrightarrow> SN t') \<Longrightarrow> SN t"
lemma SN_preserved:
- assumes a: "SN t1" "t1\<longrightarrow>\<^isub>\<beta> t2"
+ assumes a: "SN t1" "t1\<longrightarrow>\<^sub>\<beta> t2"
shows "SN t2"
using a
by (cases) (auto)
@@ -187,8 +187,8 @@
assumes a: "SN a"
and b: "SN b"
and hyp: "\<And>x z.
- \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> SN y; \<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z;
- \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> SN u; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
+ \<lbrakk>\<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> SN y; \<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> P y z;
+ \<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> SN u; \<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
shows "P a b"
proof -
from a
@@ -215,7 +215,7 @@
lemma double_SN[consumes 2]:
assumes a: "SN a"
and b: "SN b"
- and c: "\<And>x z. \<lbrakk>\<And>y. x \<longrightarrow>\<^isub>\<beta> y \<Longrightarrow> P y z; \<And>u. z \<longrightarrow>\<^isub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
+ and c: "\<And>x z. \<lbrakk>\<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> P y z; \<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z"
shows "P a b"
using a b c
apply(rule_tac double_SN_aux)
@@ -276,10 +276,10 @@
"CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)"
definition "CR2" :: "ty \<Rightarrow> bool" where
- "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>"
+ "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^sub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>"
definition "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool" where
- "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow> t'\<in>RED \<tau>"
+ "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^sub>\<beta> t' \<longrightarrow> t'\<in>RED \<tau>"
definition "CR3" :: "ty \<Rightarrow> bool" where
"CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>"
@@ -305,23 +305,23 @@
proof (induct)
fix u
assume as: "u\<in>RED \<tau>"
- assume ih: " \<And>u'. \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u' \<in> RED \<tau>\<rbrakk> \<Longrightarrow> App t u' \<in> RED \<sigma>"
+ assume ih: " \<And>u'. \<lbrakk>u \<longrightarrow>\<^sub>\<beta> u'; u' \<in> RED \<tau>\<rbrakk> \<Longrightarrow> App t u' \<in> RED \<sigma>"
have "NEUT (App t u)" using c1 by (auto simp add: NEUT_def)
moreover
have "CR3_RED (App t u) \<sigma>" unfolding CR3_RED_def
proof (intro strip)
fix r
- assume red: "App t u \<longrightarrow>\<^isub>\<beta> r"
+ assume red: "App t u \<longrightarrow>\<^sub>\<beta> r"
moreover
- { assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App t' u"
- then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App t' u" by blast
+ { assume "\<exists>t'. t \<longrightarrow>\<^sub>\<beta> t' \<and> r = App t' u"
+ then obtain t' where a1: "t \<longrightarrow>\<^sub>\<beta> t'" and a2: "r = App t' u" by blast
have "t'\<in>RED (\<tau>\<rightarrow>\<sigma>)" using c4 a1 by (simp add: CR3_RED_def)
then have "App t' u\<in>RED \<sigma>" using as by simp
then have "r\<in>RED \<sigma>" using a2 by simp
}
moreover
- { assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App t u'"
- then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App t u'" by blast
+ { assume "\<exists>u'. u \<longrightarrow>\<^sub>\<beta> u' \<and> r = App t u'"
+ then obtain u' where b1: "u \<longrightarrow>\<^sub>\<beta> u'" and b2: "r = App t u'" by blast
have "u'\<in>RED \<tau>" using as b1 c2 by (auto simp add: CR2_def)
with ih have "App t u' \<in> RED \<sigma>" using b1 by simp
then have "r\<in>RED \<sigma>" using b2 by simp
@@ -423,17 +423,17 @@
show "App (Lam [x].t) u \<in> RED \<sigma>" using b1 b3 b2 asm
proof(induct t u rule: double_SN)
fix t u
- assume ih1: "\<And>t'. \<lbrakk>t \<longrightarrow>\<^isub>\<beta> t'; u\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t'[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t') u \<in> RED \<sigma>"
- assume ih2: "\<And>u'. \<lbrakk>u \<longrightarrow>\<^isub>\<beta> u'; u'\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t) u' \<in> RED \<sigma>"
+ assume ih1: "\<And>t'. \<lbrakk>t \<longrightarrow>\<^sub>\<beta> t'; u\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t'[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t') u \<in> RED \<sigma>"
+ assume ih2: "\<And>u'. \<lbrakk>u \<longrightarrow>\<^sub>\<beta> u'; u'\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t) u' \<in> RED \<sigma>"
assume as1: "u \<in> RED \<tau>"
assume as2: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>"
have "CR3_RED (App (Lam [x].t) u) \<sigma>" unfolding CR3_RED_def
proof(intro strip)
fix r
- assume red: "App (Lam [x].t) u \<longrightarrow>\<^isub>\<beta> r"
+ assume red: "App (Lam [x].t) u \<longrightarrow>\<^sub>\<beta> r"
moreover
- { assume "\<exists>t'. t \<longrightarrow>\<^isub>\<beta> t' \<and> r = App (Lam [x].t') u"
- then obtain t' where a1: "t \<longrightarrow>\<^isub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast
+ { assume "\<exists>t'. t \<longrightarrow>\<^sub>\<beta> t' \<and> r = App (Lam [x].t') u"
+ then obtain t' where a1: "t \<longrightarrow>\<^sub>\<beta> t'" and a2: "r = App (Lam [x].t') u" by blast
have "App (Lam [x].t') u\<in>RED \<sigma>" using ih1 a1 as1 as2
apply(auto)
apply(drule_tac x="t'" in meta_spec)
@@ -455,8 +455,8 @@
then have "r\<in>RED \<sigma>" using a2 by simp
}
moreover
- { assume "\<exists>u'. u \<longrightarrow>\<^isub>\<beta> u' \<and> r = App (Lam [x].t) u'"
- then obtain u' where b1: "u \<longrightarrow>\<^isub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast
+ { assume "\<exists>u'. u \<longrightarrow>\<^sub>\<beta> u' \<and> r = App (Lam [x].t) u'"
+ then obtain u' where b1: "u \<longrightarrow>\<^sub>\<beta> u'" and b2: "r = App (Lam [x].t) u'" by blast
have "App (Lam [x].t) u'\<in>RED \<sigma>" using ih2 b1 as1 as2
apply(auto)
apply(drule_tac x="u'" in meta_spec)
--- a/src/HOL/Nominal/Examples/SOS.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/SOS.thy Tue Aug 13 16:25:47 2013 +0200
@@ -65,7 +65,7 @@
psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm" ("_<_>" [95,95] 105)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
-| "\<theta><(App e\<^isub>1 e\<^isub>2)> = App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>)"
+| "\<theta><(App e\<^sub>1 e\<^sub>2)> = App (\<theta><e\<^sub>1>) (\<theta><e\<^sub>2>)"
| "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].e)> = Lam [x].(\<theta><e>)"
apply(finite_guess)+
apply(rule TrueI)+
@@ -102,7 +102,7 @@
lemma subst[simp]:
shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
- and "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])"
+ and "(App t\<^sub>1 t\<^sub>2)[y::=t'] = App (t\<^sub>1[y::=t']) (t\<^sub>2[y::=t'])"
and "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
by (simp_all add: fresh_list_cons fresh_list_nil)
@@ -164,8 +164,8 @@
typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
where
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> e\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2"
-| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> e : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T\<^isub>1\<rightarrow>T\<^isub>2"
+| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> e\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^sub>1 e\<^sub>2 : T\<^sub>2"
+| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^sub>1)#\<Gamma> \<turnstile> e : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T\<^sub>1\<rightarrow>T\<^sub>2"
equivariance typing
@@ -186,7 +186,7 @@
lemma t_Lam_elim:
assumes a: "\<Gamma> \<turnstile> Lam [x].t : T" "x\<sharp>\<Gamma>"
- obtains T\<^isub>1 and T\<^isub>2 where "(x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2" and "T=T\<^isub>1\<rightarrow>T\<^isub>2"
+ obtains T\<^sub>1 and T\<^sub>2 where "(x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2" and "T=T\<^sub>1\<rightarrow>T\<^sub>2"
using a
by (cases rule: typing.strong_cases [where x="x"])
(auto simp add: abs_fresh fresh_ty alpha trm.inject)
@@ -194,24 +194,24 @@
abbreviation
"sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [55,55] 55)
where
- "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^isub>2"
+ "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^sub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^sub>2"
lemma weakening:
- fixes \<Gamma>\<^isub>1 \<Gamma>\<^isub>2::"(name\<times>ty) list"
- assumes "\<Gamma>\<^isub>1 \<turnstile> e: T" and "valid \<Gamma>\<^isub>2" and "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2"
- shows "\<Gamma>\<^isub>2 \<turnstile> e: T"
+ fixes \<Gamma>\<^sub>1 \<Gamma>\<^sub>2::"(name\<times>ty) list"
+ assumes "\<Gamma>\<^sub>1 \<turnstile> e: T" and "valid \<Gamma>\<^sub>2" and "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2"
+ shows "\<Gamma>\<^sub>2 \<turnstile> e: T"
using assms
-proof(nominal_induct \<Gamma>\<^isub>1 e T avoiding: \<Gamma>\<^isub>2 rule: typing.strong_induct)
- case (t_Lam x \<Gamma>\<^isub>1 T\<^isub>1 t T\<^isub>2 \<Gamma>\<^isub>2)
- have vc: "x\<sharp>\<Gamma>\<^isub>2" by fact
- have ih: "\<lbrakk>valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2); (x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2\<rbrakk> \<Longrightarrow> (x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" by fact
- have "valid \<Gamma>\<^isub>2" by fact
- then have "valid ((x,T\<^isub>1)#\<Gamma>\<^isub>2)" using vc by auto
+proof(nominal_induct \<Gamma>\<^sub>1 e T avoiding: \<Gamma>\<^sub>2 rule: typing.strong_induct)
+ case (t_Lam x \<Gamma>\<^sub>1 T\<^sub>1 t T\<^sub>2 \<Gamma>\<^sub>2)
+ have vc: "x\<sharp>\<Gamma>\<^sub>2" by fact
+ have ih: "\<lbrakk>valid ((x,T\<^sub>1)#\<Gamma>\<^sub>2); (x,T\<^sub>1)#\<Gamma>\<^sub>1 \<subseteq> (x,T\<^sub>1)#\<Gamma>\<^sub>2\<rbrakk> \<Longrightarrow> (x,T\<^sub>1)#\<Gamma>\<^sub>2 \<turnstile> t : T\<^sub>2" by fact
+ have "valid \<Gamma>\<^sub>2" by fact
+ then have "valid ((x,T\<^sub>1)#\<Gamma>\<^sub>2)" using vc by auto
moreover
- have "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2" by fact
- then have "(x,T\<^isub>1)#\<Gamma>\<^isub>1 \<subseteq> (x,T\<^isub>1)#\<Gamma>\<^isub>2" by simp
- ultimately have "(x,T\<^isub>1)#\<Gamma>\<^isub>2 \<turnstile> t : T\<^isub>2" using ih by simp
- with vc show "\<Gamma>\<^isub>2 \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
+ have "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2" by fact
+ then have "(x,T\<^sub>1)#\<Gamma>\<^sub>1 \<subseteq> (x,T\<^sub>1)#\<Gamma>\<^sub>2" by simp
+ ultimately have "(x,T\<^sub>1)#\<Gamma>\<^sub>2 \<turnstile> t : T\<^sub>2" using ih by simp
+ with vc show "\<Gamma>\<^sub>2 \<turnstile> Lam [x].t : T\<^sub>1\<rightarrow>T\<^sub>2" by auto
qed (auto)
lemma type_substitutivity_aux:
@@ -254,7 +254,7 @@
lemma not_val_App[simp]:
shows
- "\<not> val (App e\<^isub>1 e\<^isub>2)"
+ "\<not> val (App e\<^sub>1 e\<^sub>2)"
"\<not> val (Var x)"
by (auto elim: val.cases)
@@ -264,7 +264,7 @@
big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
where
b_Lam[intro]: "Lam [x].e \<Down> Lam [x].e"
-| b_App[intro]: "\<lbrakk>x\<sharp>(e\<^isub>1,e\<^isub>2,e'); e\<^isub>1\<Down>Lam [x].e; e\<^isub>2\<Down>e\<^isub>2'; e[x::=e\<^isub>2']\<Down>e'\<rbrakk> \<Longrightarrow> App e\<^isub>1 e\<^isub>2 \<Down> e'"
+| b_App[intro]: "\<lbrakk>x\<sharp>(e\<^sub>1,e\<^sub>2,e'); e\<^sub>1\<Down>Lam [x].e; e\<^sub>2\<Down>e\<^sub>2'; e[x::=e\<^sub>2']\<Down>e'\<rbrakk> \<Longrightarrow> App e\<^sub>1 e\<^sub>2 \<Down> e'"
equivariance big
@@ -278,8 +278,8 @@
using a by (induct) (auto simp add: abs_fresh fresh_subst)
lemma b_App_elim:
- assumes a: "App e\<^isub>1 e\<^isub>2 \<Down> e'" "x\<sharp>(e\<^isub>1,e\<^isub>2,e')"
- obtains f\<^isub>1 and f\<^isub>2 where "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" "e\<^isub>2 \<Down> f\<^isub>2" "f\<^isub>1[x::=f\<^isub>2] \<Down> e'"
+ assumes a: "App e\<^sub>1 e\<^sub>2 \<Down> e'" "x\<sharp>(e\<^sub>1,e\<^sub>2,e')"
+ obtains f\<^sub>1 and f\<^sub>2 where "e\<^sub>1 \<Down> Lam [x]. f\<^sub>1" "e\<^sub>2 \<Down> f\<^sub>2" "f\<^sub>1[x::=f\<^sub>2] \<Down> e'"
using a
by (cases rule: big.strong_cases[where x="x" and xa="x"])
(auto simp add: trm.inject)
@@ -289,20 +289,20 @@
shows "\<Gamma> \<turnstile> e' : T"
using a b
proof (nominal_induct avoiding: \<Gamma> arbitrary: T rule: big.strong_induct)
- case (b_App x e\<^isub>1 e\<^isub>2 e' e e\<^isub>2' \<Gamma> T)
+ case (b_App x e\<^sub>1 e\<^sub>2 e' e e\<^sub>2' \<Gamma> T)
have vc: "x\<sharp>\<Gamma>" by fact
- have "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
- then obtain T' where a1: "\<Gamma> \<turnstile> e\<^isub>1 : T'\<rightarrow>T" and a2: "\<Gamma> \<turnstile> e\<^isub>2 : T'"
+ have "\<Gamma> \<turnstile> App e\<^sub>1 e\<^sub>2 : T" by fact
+ then obtain T' where a1: "\<Gamma> \<turnstile> e\<^sub>1 : T'\<rightarrow>T" and a2: "\<Gamma> \<turnstile> e\<^sub>2 : T'"
by (cases) (auto simp add: trm.inject)
- have ih1: "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T' \<rightarrow> T" by fact
- have ih2: "\<Gamma> \<turnstile> e\<^isub>2 : T' \<Longrightarrow> \<Gamma> \<turnstile> e\<^isub>2' : T'" by fact
- have ih3: "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T \<Longrightarrow> \<Gamma> \<turnstile> e' : T" by fact
+ have ih1: "\<Gamma> \<turnstile> e\<^sub>1 : T' \<rightarrow> T \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].e : T' \<rightarrow> T" by fact
+ have ih2: "\<Gamma> \<turnstile> e\<^sub>2 : T' \<Longrightarrow> \<Gamma> \<turnstile> e\<^sub>2' : T'" by fact
+ have ih3: "\<Gamma> \<turnstile> e[x::=e\<^sub>2'] : T \<Longrightarrow> \<Gamma> \<turnstile> e' : T" by fact
have "\<Gamma> \<turnstile> Lam [x].e : T'\<rightarrow>T" using ih1 a1 by simp
then have "((x,T')#\<Gamma>) \<turnstile> e : T" using vc
by (auto elim: t_Lam_elim simp add: ty.inject)
moreover
- have "\<Gamma> \<turnstile> e\<^isub>2': T'" using ih2 a2 by simp
- ultimately have "\<Gamma> \<turnstile> e[x::=e\<^isub>2'] : T" by (simp add: type_substitutivity)
+ have "\<Gamma> \<turnstile> e\<^sub>2': T'" using ih2 a2 by simp
+ ultimately have "\<Gamma> \<turnstile> e[x::=e\<^sub>2'] : T" by (simp add: type_substitutivity)
thus "\<Gamma> \<turnstile> e' : T" using ih3 by simp
qed (blast)
@@ -314,31 +314,31 @@
(force elim: t_App_elim t_Lam_elim simp add: ty.inject type_substitutivity)+
lemma unicity_of_evaluation:
- assumes a: "e \<Down> e\<^isub>1"
- and b: "e \<Down> e\<^isub>2"
- shows "e\<^isub>1 = e\<^isub>2"
+ assumes a: "e \<Down> e\<^sub>1"
+ and b: "e \<Down> e\<^sub>2"
+ shows "e\<^sub>1 = e\<^sub>2"
using a b
-proof (nominal_induct e e\<^isub>1 avoiding: e\<^isub>2 rule: big.strong_induct)
- case (b_Lam x e t\<^isub>2)
- have "Lam [x].e \<Down> t\<^isub>2" by fact
- thus "Lam [x].e = t\<^isub>2" by cases (simp_all add: trm.inject)
+proof (nominal_induct e e\<^sub>1 avoiding: e\<^sub>2 rule: big.strong_induct)
+ case (b_Lam x e t\<^sub>2)
+ have "Lam [x].e \<Down> t\<^sub>2" by fact
+ thus "Lam [x].e = t\<^sub>2" by cases (simp_all add: trm.inject)
next
- case (b_App x e\<^isub>1 e\<^isub>2 e' e\<^isub>1' e\<^isub>2' t\<^isub>2)
- have ih1: "\<And>t. e\<^isub>1 \<Down> t \<Longrightarrow> Lam [x].e\<^isub>1' = t" by fact
- have ih2:"\<And>t. e\<^isub>2 \<Down> t \<Longrightarrow> e\<^isub>2' = t" by fact
- have ih3: "\<And>t. e\<^isub>1'[x::=e\<^isub>2'] \<Down> t \<Longrightarrow> e' = t" by fact
- have app: "App e\<^isub>1 e\<^isub>2 \<Down> t\<^isub>2" by fact
- have vc: "x\<sharp>e\<^isub>1" "x\<sharp>e\<^isub>2" "x\<sharp>t\<^isub>2" by fact+
- then have "x\<sharp>App e\<^isub>1 e\<^isub>2" by auto
- from app vc obtain f\<^isub>1 f\<^isub>2 where x1: "e\<^isub>1 \<Down> Lam [x]. f\<^isub>1" and x2: "e\<^isub>2 \<Down> f\<^isub>2" and x3: "f\<^isub>1[x::=f\<^isub>2] \<Down> t\<^isub>2"
+ case (b_App x e\<^sub>1 e\<^sub>2 e' e\<^sub>1' e\<^sub>2' t\<^sub>2)
+ have ih1: "\<And>t. e\<^sub>1 \<Down> t \<Longrightarrow> Lam [x].e\<^sub>1' = t" by fact
+ have ih2:"\<And>t. e\<^sub>2 \<Down> t \<Longrightarrow> e\<^sub>2' = t" by fact
+ have ih3: "\<And>t. e\<^sub>1'[x::=e\<^sub>2'] \<Down> t \<Longrightarrow> e' = t" by fact
+ have app: "App e\<^sub>1 e\<^sub>2 \<Down> t\<^sub>2" by fact
+ have vc: "x\<sharp>e\<^sub>1" "x\<sharp>e\<^sub>2" "x\<sharp>t\<^sub>2" by fact+
+ then have "x\<sharp>App e\<^sub>1 e\<^sub>2" by auto
+ from app vc obtain f\<^sub>1 f\<^sub>2 where x1: "e\<^sub>1 \<Down> Lam [x]. f\<^sub>1" and x2: "e\<^sub>2 \<Down> f\<^sub>2" and x3: "f\<^sub>1[x::=f\<^sub>2] \<Down> t\<^sub>2"
by (auto elim!: b_App_elim)
- then have "Lam [x]. f\<^isub>1 = Lam [x]. e\<^isub>1'" using ih1 by simp
+ then have "Lam [x]. f\<^sub>1 = Lam [x]. e\<^sub>1'" using ih1 by simp
then
- have "f\<^isub>1 = e\<^isub>1'" by (auto simp add: trm.inject alpha)
+ have "f\<^sub>1 = e\<^sub>1'" by (auto simp add: trm.inject alpha)
moreover
- have "f\<^isub>2 = e\<^isub>2'" using x2 ih2 by simp
- ultimately have "e\<^isub>1'[x::=e\<^isub>2'] \<Down> t\<^isub>2" using x3 by simp
- thus "e' = t\<^isub>2" using ih3 by simp
+ have "f\<^sub>2 = e\<^sub>2'" using x2 ih2 by simp
+ ultimately have "e\<^sub>1'[x::=e\<^sub>2'] \<Down> t\<^sub>2" using x3 by simp
+ thus "e' = t\<^sub>2" using ih3 by simp
qed
lemma reduces_evaluates_to_values:
@@ -352,7 +352,7 @@
V :: "ty \<Rightarrow> trm set"
where
"V (TVar x) = {e. val e}"
-| "V (T\<^isub>1 \<rightarrow> T\<^isub>2) = {Lam [x].e | x e. \<forall> v \<in> (V T\<^isub>1). \<exists> v'. e[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2}"
+| "V (T\<^sub>1 \<rightarrow> T\<^sub>2) = {Lam [x].e | x e. \<forall> v \<in> (V T\<^sub>1). \<exists> v'. e[x::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2}"
by (rule TrueI)+
lemma V_eqvt:
@@ -377,14 +377,14 @@
done
lemma V_arrow_elim_weak:
- assumes h:"u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)"
- obtains a t where "u = Lam [a].t" and "\<forall> v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
+ assumes h:"u \<in> V (T\<^sub>1 \<rightarrow> T\<^sub>2)"
+ obtains a t where "u = Lam [a].t" and "\<forall> v \<in> (V T\<^sub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2"
using h by (auto)
lemma V_arrow_elim_strong:
fixes c::"'a::fs_name"
- assumes h: "u \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)"
- obtains a t where "a\<sharp>c" "u = Lam [a].t" "\<forall>v \<in> (V T\<^isub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
+ assumes h: "u \<in> V (T\<^sub>1 \<rightarrow> T\<^sub>2)"
+ obtains a t where "a\<sharp>c" "u = Lam [a].t" "\<forall>v \<in> (V T\<^sub>1). \<exists> v'. t[a::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2"
using h
apply -
apply(erule V_arrow_elim_weak)
@@ -489,48 +489,48 @@
shows "\<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T"
using h2 h1
proof(nominal_induct e avoiding: \<Gamma> \<theta> arbitrary: T rule: trm.strong_induct)
- case (App e\<^isub>1 e\<^isub>2 \<Gamma> \<theta> T)
- have ih\<^isub>1: "\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>1 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>1> \<Down> v \<and> v \<in> V T" by fact
- have ih\<^isub>2: "\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^isub>2 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^isub>2> \<Down> v \<and> v \<in> V T" by fact
- have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
- have as\<^isub>2: "\<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T" by fact
- then obtain T' where "\<Gamma> \<turnstile> e\<^isub>1 : T' \<rightarrow> T" and "\<Gamma> \<turnstile> e\<^isub>2 : T'" by (auto elim: t_App_elim)
- then obtain v\<^isub>1 v\<^isub>2 where "(i)": "\<theta><e\<^isub>1> \<Down> v\<^isub>1" "v\<^isub>1 \<in> V (T' \<rightarrow> T)"
- and "(ii)": "\<theta><e\<^isub>2> \<Down> v\<^isub>2" "v\<^isub>2 \<in> V T'" using ih\<^isub>1 ih\<^isub>2 as\<^isub>1 by blast
+ case (App e\<^sub>1 e\<^sub>2 \<Gamma> \<theta> T)
+ have ih\<^sub>1: "\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^sub>1 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^sub>1> \<Down> v \<and> v \<in> V T" by fact
+ have ih\<^sub>2: "\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e\<^sub>2 : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e\<^sub>2> \<Down> v \<and> v \<in> V T" by fact
+ have as\<^sub>1: "\<theta> Vcloses \<Gamma>" by fact
+ have as\<^sub>2: "\<Gamma> \<turnstile> App e\<^sub>1 e\<^sub>2 : T" by fact
+ then obtain T' where "\<Gamma> \<turnstile> e\<^sub>1 : T' \<rightarrow> T" and "\<Gamma> \<turnstile> e\<^sub>2 : T'" by (auto elim: t_App_elim)
+ then obtain v\<^sub>1 v\<^sub>2 where "(i)": "\<theta><e\<^sub>1> \<Down> v\<^sub>1" "v\<^sub>1 \<in> V (T' \<rightarrow> T)"
+ and "(ii)": "\<theta><e\<^sub>2> \<Down> v\<^sub>2" "v\<^sub>2 \<in> V T'" using ih\<^sub>1 ih\<^sub>2 as\<^sub>1 by blast
from "(i)" obtain x e'
- where "v\<^isub>1 = Lam [x].e'"
+ where "v\<^sub>1 = Lam [x].e'"
and "(iii)": "(\<forall>v \<in> (V T').\<exists> v'. e'[x::=v] \<Down> v' \<and> v' \<in> V T)"
- and "(iv)": "\<theta><e\<^isub>1> \<Down> Lam [x].e'"
- and fr: "x\<sharp>(\<theta>,e\<^isub>1,e\<^isub>2)" by (blast elim: V_arrow_elim_strong)
- from fr have fr\<^isub>1: "x\<sharp>\<theta><e\<^isub>1>" and fr\<^isub>2: "x\<sharp>\<theta><e\<^isub>2>" by (simp_all add: fresh_psubst)
- from "(ii)" "(iii)" obtain v\<^isub>3 where "(v)": "e'[x::=v\<^isub>2] \<Down> v\<^isub>3" "v\<^isub>3 \<in> V T" by auto
- from fr\<^isub>2 "(ii)" have "x\<sharp>v\<^isub>2" by (simp add: big_preserves_fresh)
- then have "x\<sharp>e'[x::=v\<^isub>2]" by (simp add: fresh_subst)
- then have fr\<^isub>3: "x\<sharp>v\<^isub>3" using "(v)" by (simp add: big_preserves_fresh)
- from fr\<^isub>1 fr\<^isub>2 fr\<^isub>3 have "x\<sharp>(\<theta><e\<^isub>1>,\<theta><e\<^isub>2>,v\<^isub>3)" by simp
- with "(iv)" "(ii)" "(v)" have "App (\<theta><e\<^isub>1>) (\<theta><e\<^isub>2>) \<Down> v\<^isub>3" by auto
- then show "\<exists>v. \<theta><App e\<^isub>1 e\<^isub>2> \<Down> v \<and> v \<in> V T" using "(v)" by auto
+ and "(iv)": "\<theta><e\<^sub>1> \<Down> Lam [x].e'"
+ and fr: "x\<sharp>(\<theta>,e\<^sub>1,e\<^sub>2)" by (blast elim: V_arrow_elim_strong)
+ from fr have fr\<^sub>1: "x\<sharp>\<theta><e\<^sub>1>" and fr\<^sub>2: "x\<sharp>\<theta><e\<^sub>2>" by (simp_all add: fresh_psubst)
+ from "(ii)" "(iii)" obtain v\<^sub>3 where "(v)": "e'[x::=v\<^sub>2] \<Down> v\<^sub>3" "v\<^sub>3 \<in> V T" by auto
+ from fr\<^sub>2 "(ii)" have "x\<sharp>v\<^sub>2" by (simp add: big_preserves_fresh)
+ then have "x\<sharp>e'[x::=v\<^sub>2]" by (simp add: fresh_subst)
+ then have fr\<^sub>3: "x\<sharp>v\<^sub>3" using "(v)" by (simp add: big_preserves_fresh)
+ from fr\<^sub>1 fr\<^sub>2 fr\<^sub>3 have "x\<sharp>(\<theta><e\<^sub>1>,\<theta><e\<^sub>2>,v\<^sub>3)" by simp
+ with "(iv)" "(ii)" "(v)" have "App (\<theta><e\<^sub>1>) (\<theta><e\<^sub>2>) \<Down> v\<^sub>3" by auto
+ then show "\<exists>v. \<theta><App e\<^sub>1 e\<^sub>2> \<Down> v \<and> v \<in> V T" using "(v)" by auto
next
case (Lam x e \<Gamma> \<theta> T)
have ih:"\<And>\<theta> \<Gamma> T. \<lbrakk>\<theta> Vcloses \<Gamma>; \<Gamma> \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>v. \<theta><e> \<Down> v \<and> v \<in> V T" by fact
- have as\<^isub>1: "\<theta> Vcloses \<Gamma>" by fact
- have as\<^isub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
+ have as\<^sub>1: "\<theta> Vcloses \<Gamma>" by fact
+ have as\<^sub>2: "\<Gamma> \<turnstile> Lam [x].e : T" by fact
have fs: "x\<sharp>\<Gamma>" "x\<sharp>\<theta>" by fact+
- from as\<^isub>2 fs obtain T\<^isub>1 T\<^isub>2
- where "(i)": "(x,T\<^isub>1)#\<Gamma> \<turnstile> e:T\<^isub>2" and "(ii)": "T = T\<^isub>1 \<rightarrow> T\<^isub>2" using fs
+ from as\<^sub>2 fs obtain T\<^sub>1 T\<^sub>2
+ where "(i)": "(x,T\<^sub>1)#\<Gamma> \<turnstile> e:T\<^sub>2" and "(ii)": "T = T\<^sub>1 \<rightarrow> T\<^sub>2" using fs
by (auto elim: t_Lam_elim)
- from "(i)" have "(iii)": "valid ((x,T\<^isub>1)#\<Gamma>)" by (simp add: typing_implies_valid)
- have "\<forall>v \<in> (V T\<^isub>1). \<exists>v'. (\<theta><e>)[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2"
+ from "(i)" have "(iii)": "valid ((x,T\<^sub>1)#\<Gamma>)" by (simp add: typing_implies_valid)
+ have "\<forall>v \<in> (V T\<^sub>1). \<exists>v'. (\<theta><e>)[x::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2"
proof
fix v
- assume "v \<in> (V T\<^isub>1)"
- with "(iii)" as\<^isub>1 have "(x,v)#\<theta> Vcloses (x,T\<^isub>1)#\<Gamma>" using monotonicity by auto
- with ih "(i)" obtain v' where "((x,v)#\<theta>)<e> \<Down> v' \<and> v' \<in> V T\<^isub>2" by blast
- then have "\<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" using fs by (simp add: psubst_subst_psubst)
- then show "\<exists>v'. \<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^isub>2" by auto
+ assume "v \<in> (V T\<^sub>1)"
+ with "(iii)" as\<^sub>1 have "(x,v)#\<theta> Vcloses (x,T\<^sub>1)#\<Gamma>" using monotonicity by auto
+ with ih "(i)" obtain v' where "((x,v)#\<theta>)<e> \<Down> v' \<and> v' \<in> V T\<^sub>2" by blast
+ then have "\<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2" using fs by (simp add: psubst_subst_psubst)
+ then show "\<exists>v'. \<theta><e>[x::=v] \<Down> v' \<and> v' \<in> V T\<^sub>2" by auto
qed
- then have "Lam[x].\<theta><e> \<in> V (T\<^isub>1 \<rightarrow> T\<^isub>2)" by auto
- then have "\<theta><Lam [x].e> \<Down> Lam [x].\<theta><e> \<and> Lam [x].\<theta><e> \<in> V (T\<^isub>1\<rightarrow>T\<^isub>2)" using fs by auto
+ then have "Lam[x].\<theta><e> \<in> V (T\<^sub>1 \<rightarrow> T\<^sub>2)" by auto
+ then have "\<theta><Lam [x].e> \<Down> Lam [x].\<theta><e> \<and> Lam [x].\<theta><e> \<in> V (T\<^sub>1\<rightarrow>T\<^sub>2)" using fs by auto
thus "\<exists>v. \<theta><Lam [x].e> \<Down> v \<and> v \<in> V T" using "(ii)" by auto
next
case (Var x \<Gamma> \<theta> T)
--- a/src/HOL/Nominal/Examples/Standardization.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Standardization.thy Tue Aug 13 16:25:47 2013 +0200
@@ -45,7 +45,7 @@
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [300, 0, 0] 300)
where
subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))"
-| subst_App: "(t\<^isub>1 \<degree> t\<^isub>2)[y::=s] = t\<^isub>1[y::=s] \<degree> t\<^isub>2[y::=s]"
+| subst_App: "(t\<^sub>1 \<degree> t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \<degree> t\<^sub>2[y::=s]"
| subst_Lam: "x \<sharp> (y, s) \<Longrightarrow> (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))"
apply(finite_guess)+
apply(rule TrueI)+
--- a/src/HOL/Nominal/Examples/Type_Preservation.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Tue Aug 13 16:25:47 2013 +0200
@@ -23,7 +23,7 @@
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]")
where
"(Var x)[y::=s] = (if x=y then s else (Var x))"
-| "(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])"
+| "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])"
| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])"
apply(finite_guess)+
apply(rule TrueI)+
@@ -63,7 +63,7 @@
abbreviation
"sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" ("_ \<subseteq> _")
where
- "\<Gamma>\<^isub>1 \<subseteq> \<Gamma>\<^isub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^isub>1 \<longrightarrow> x \<in> set \<Gamma>\<^isub>2"
+ "\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
text {* Validity of Typing Contexts *}
@@ -105,8 +105,8 @@
typing :: "ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
where
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t\<^isub>1 t\<^isub>2 : T\<^isub>2"
-| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2"
+| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t\<^sub>1 t\<^sub>2 : T\<^sub>2"
+| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^sub>1\<rightarrow>T\<^sub>2"
equivariance typing
nominal_inductive typing
@@ -115,7 +115,7 @@
lemma t_Lam_inversion[dest]:
assumes ty: "\<Gamma> \<turnstile> Lam [x].t : T"
and fc: "x\<sharp>\<Gamma>"
- shows "\<exists>T\<^isub>1 T\<^isub>2. T = T\<^isub>1 \<rightarrow> T\<^isub>2 \<and> (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2"
+ shows "\<exists>T\<^sub>1 T\<^sub>2. T = T\<^sub>1 \<rightarrow> T\<^sub>2 \<and> (x,T\<^sub>1)#\<Gamma> \<turnstile> t : T\<^sub>2"
using ty fc
by (cases rule: typing.strong_cases)
(auto simp add: alpha lam.inject abs_fresh ty_fresh)
@@ -169,12 +169,12 @@
text {* Beta Reduction *}
inductive
- "beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _")
+ "beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _")
where
- b1[intro]: "t1 \<longrightarrow>\<^isub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^isub>\<beta> App t2 s"
-| b2[intro]: "s1 \<longrightarrow>\<^isub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^isub>\<beta> App t s2"
-| b3[intro]: "t1 \<longrightarrow>\<^isub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^isub>\<beta> Lam [x].t2"
-| b4[intro]: "x\<sharp>s \<Longrightarrow> App (Lam [x].t) s \<longrightarrow>\<^isub>\<beta> t[x::=s]"
+ b1[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^sub>\<beta> App t2 s"
+| b2[intro]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
+| b3[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>\<beta> Lam [x].t2"
+| b4[intro]: "x\<sharp>s \<Longrightarrow> App (Lam [x].t) s \<longrightarrow>\<^sub>\<beta> t[x::=s]"
equivariance beta
nominal_inductive beta
@@ -182,7 +182,7 @@
theorem type_preservation:
- assumes a: "t \<longrightarrow>\<^isub>\<beta> t'"
+ assumes a: "t \<longrightarrow>\<^sub>\<beta> t'"
and b: "\<Gamma> \<turnstile> t : T"
shows "\<Gamma> \<turnstile> t' : T"
using a b
@@ -216,7 +216,7 @@
theorem type_preservation_automatic:
- assumes a: "t \<longrightarrow>\<^isub>\<beta> t'"
+ assumes a: "t \<longrightarrow>\<^sub>\<beta> t'"
and b: "\<Gamma> \<turnstile> t : T"
shows "\<Gamma> \<turnstile> t' : T"
using a b
--- a/src/HOL/Nominal/Examples/W.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Nominal/Examples/W.thy Tue Aug 13 16:25:47 2013 +0200
@@ -210,7 +210,7 @@
psubst_ty
where
"\<theta><TVar X> = lookup \<theta> X"
-| "\<theta><T\<^isub>1 \<rightarrow> T\<^isub>2> = (\<theta><T\<^isub>1>) \<rightarrow> (\<theta><T\<^isub>2>)"
+| "\<theta><T\<^sub>1 \<rightarrow> T\<^sub>2> = (\<theta><T\<^sub>1>) \<rightarrow> (\<theta><T\<^sub>2>)"
by (rule TrueI)+
end
@@ -357,10 +357,10 @@
typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
where
T_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,S)\<in>set \<Gamma>; T \<prec> S\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| T_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t\<^isub>1 t\<^isub>2 : T\<^isub>2"
-| T_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,Ty T\<^isub>1)#\<Gamma>) \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2"
-| T_LET[intro]: "\<lbrakk>x\<sharp>\<Gamma>; \<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1; ((x,close \<Gamma> T\<^isub>1)#\<Gamma>) \<turnstile> t\<^isub>2 : T\<^isub>2; set (ftv T\<^isub>1 - ftv \<Gamma>) \<sharp>* T\<^isub>2\<rbrakk>
- \<Longrightarrow> \<Gamma> \<turnstile> Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2"
+| T_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t\<^sub>1 t\<^sub>2 : T\<^sub>2"
+| T_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,Ty T\<^sub>1)#\<Gamma>) \<turnstile> t : T\<^sub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^sub>1\<rightarrow>T\<^sub>2"
+| T_LET[intro]: "\<lbrakk>x\<sharp>\<Gamma>; \<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1; ((x,close \<Gamma> T\<^sub>1)#\<Gamma>) \<turnstile> t\<^sub>2 : T\<^sub>2; set (ftv T\<^sub>1 - ftv \<Gamma>) \<sharp>* T\<^sub>2\<rbrakk>
+ \<Longrightarrow> \<Gamma> \<turnstile> Let x be t\<^sub>1 in t\<^sub>2 : T\<^sub>2"
equivariance typing[tvar]
@@ -411,7 +411,7 @@
(simp_all add: supp_list_nil supp_list_cons supp_atm)
nominal_inductive2 typing
- avoids T_LET: "set (ftv T\<^isub>1 - ftv \<Gamma>)"
+ avoids T_LET: "set (ftv T\<^sub>1 - ftv \<Gamma>)"
apply (simp add: fresh_star_def fresh_def ftv_Ctxt)
apply (simp add: fresh_star_def fresh_tvar_trm)
apply assumption
@@ -470,15 +470,15 @@
lemma better_T_LET:
assumes x: "x\<sharp>\<Gamma>"
- and t1: "\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1"
- and t2: "((x,close \<Gamma> T\<^isub>1)#\<Gamma>) \<turnstile> t\<^isub>2 : T\<^isub>2"
- shows "\<Gamma> \<turnstile> Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2"
+ and t1: "\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1"
+ and t2: "((x,close \<Gamma> T\<^sub>1)#\<Gamma>) \<turnstile> t\<^sub>2 : T\<^sub>2"
+ shows "\<Gamma> \<turnstile> Let x be t\<^sub>1 in t\<^sub>2 : T\<^sub>2"
proof -
- have fin: "finite (set (ftv T\<^isub>1 - ftv \<Gamma>))" by simp
- obtain pi where pi1: "(pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>)) \<sharp>* (T\<^isub>2, \<Gamma>)"
- and pi2: "set pi \<subseteq> set (ftv T\<^isub>1 - ftv \<Gamma>) \<times> (pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>))"
- by (rule at_set_avoiding [OF at_tvar_inst fin fs_tvar1, of "(T\<^isub>2, \<Gamma>)"])
- from pi1 have pi1': "(pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>)) \<sharp>* \<Gamma>"
+ have fin: "finite (set (ftv T\<^sub>1 - ftv \<Gamma>))" by simp
+ obtain pi where pi1: "(pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>)) \<sharp>* (T\<^sub>2, \<Gamma>)"
+ and pi2: "set pi \<subseteq> set (ftv T\<^sub>1 - ftv \<Gamma>) \<times> (pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>))"
+ by (rule at_set_avoiding [OF at_tvar_inst fin fs_tvar1, of "(T\<^sub>2, \<Gamma>)"])
+ from pi1 have pi1': "(pi \<bullet> set (ftv T\<^sub>1 - ftv \<Gamma>)) \<sharp>* \<Gamma>"
by (simp add: fresh_star_prod)
have Gamma_fresh: "\<forall>(x,y)\<in>set pi. x \<sharp> \<Gamma> \<and> y \<sharp> \<Gamma>"
apply (rule ballI)
@@ -488,7 +488,7 @@
apply (drule freshs_mem [OF _ pi1'])
apply (simp add: ftv_Ctxt [symmetric] fresh_def)
done
- have close_fresh': "\<forall>(x, y)\<in>set pi. x \<sharp> close \<Gamma> T\<^isub>1 \<and> y \<sharp> close \<Gamma> T\<^isub>1"
+ have close_fresh': "\<forall>(x, y)\<in>set pi. x \<sharp> close \<Gamma> T\<^sub>1 \<and> y \<sharp> close \<Gamma> T\<^sub>1"
apply (rule ballI)
apply (simp add: split_paired_all)
apply (drule subsetD [OF pi2])
@@ -499,15 +499,15 @@
done
note x
moreover from Gamma_fresh perm_boolI [OF t1, of pi]
- have "\<Gamma> \<turnstile> t\<^isub>1 : pi \<bullet> T\<^isub>1"
+ have "\<Gamma> \<turnstile> t\<^sub>1 : pi \<bullet> T\<^sub>1"
by (simp add: perm_fresh_fresh_aux eqvts fresh_tvar_trm)
moreover from t2 close_fresh'
- have "(x,(pi \<bullet> close \<Gamma> T\<^isub>1))#\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
+ have "(x,(pi \<bullet> close \<Gamma> T\<^sub>1))#\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
by (simp add: perm_fresh_fresh_aux)
- with Gamma_fresh have "(x,close \<Gamma> (pi \<bullet> T\<^isub>1))#\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2"
+ with Gamma_fresh have "(x,close \<Gamma> (pi \<bullet> T\<^sub>1))#\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2"
by (simp add: close_eqvt perm_fresh_fresh_aux)
moreover from pi1 Gamma_fresh
- have "set (ftv (pi \<bullet> T\<^isub>1) - ftv \<Gamma>) \<sharp>* T\<^isub>2"
+ have "set (ftv (pi \<bullet> T\<^sub>1) - ftv \<Gamma>) \<sharp>* T\<^sub>2"
by (simp only: eqvts fresh_star_prod perm_fresh_fresh_aux)
ultimately show ?thesis by (rule T_LET)
qed
--- a/src/HOL/NthRoot.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/NthRoot.thy Tue Aug 13 16:25:47 2013 +0200
@@ -349,19 +349,19 @@
lemma pos2: "0 < (2::nat)" by simp
-lemma real_sqrt_unique: "\<lbrakk>y\<twosuperior> = x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt x = y"
+lemma real_sqrt_unique: "\<lbrakk>y\<^sup>2 = x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt x = y"
unfolding sqrt_def by (rule real_root_pos_unique [OF pos2])
-lemma real_sqrt_abs [simp]: "sqrt (x\<twosuperior>) = \<bar>x\<bar>"
+lemma real_sqrt_abs [simp]: "sqrt (x\<^sup>2) = \<bar>x\<bar>"
apply (rule real_sqrt_unique)
apply (rule power2_abs)
apply (rule abs_ge_zero)
done
-lemma real_sqrt_pow2 [simp]: "0 \<le> x \<Longrightarrow> (sqrt x)\<twosuperior> = x"
+lemma real_sqrt_pow2 [simp]: "0 \<le> x \<Longrightarrow> (sqrt x)\<^sup>2 = x"
unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2])
-lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\<twosuperior> = x) = (0 \<le> x)"
+lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\<^sup>2 = x) = (0 \<le> x)"
apply (rule iffI)
apply (erule subst)
apply (rule zero_le_power2)
@@ -501,47 +501,47 @@
lemma four_x_squared:
fixes x::real
- shows "4 * x\<twosuperior> = (2 * x)\<twosuperior>"
+ shows "4 * x\<^sup>2 = (2 * x)\<^sup>2"
by (simp add: power2_eq_square)
subsection {* Square Root of Sum of Squares *}
-lemma real_sqrt_sum_squares_ge_zero: "0 \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma real_sqrt_sum_squares_ge_zero: "0 \<le> sqrt (x\<^sup>2 + y\<^sup>2)"
by simp (* TODO: delete *)
declare real_sqrt_sum_squares_ge_zero [THEN abs_of_nonneg, simp]
lemma real_sqrt_sum_squares_mult_ge_zero [simp]:
- "0 \<le> sqrt ((x\<twosuperior> + y\<twosuperior>)*(xa\<twosuperior> + ya\<twosuperior>))"
+ "0 \<le> sqrt ((x\<^sup>2 + y\<^sup>2)*(xa\<^sup>2 + ya\<^sup>2))"
by (simp add: zero_le_mult_iff)
lemma real_sqrt_sum_squares_mult_squared_eq [simp]:
- "sqrt ((x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)) ^ 2 = (x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)"
+ "sqrt ((x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)) ^ 2 = (x\<^sup>2 + y\<^sup>2) * (xa\<^sup>2 + ya\<^sup>2)"
by (simp add: zero_le_mult_iff)
-lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<twosuperior> + y\<twosuperior>) = x \<Longrightarrow> y = 0"
-by (drule_tac f = "%x. x\<twosuperior>" in arg_cong, simp)
+lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<^sup>2 + y\<^sup>2) = x \<Longrightarrow> y = 0"
+by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp)
-lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<twosuperior> + y\<twosuperior>) = y \<Longrightarrow> x = 0"
-by (drule_tac f = "%x. x\<twosuperior>" in arg_cong, simp)
+lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<^sup>2 + y\<^sup>2) = y \<Longrightarrow> x = 0"
+by (drule_tac f = "%x. x\<^sup>2" in arg_cong, simp)
-lemma real_sqrt_sum_squares_ge1 [simp]: "x \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma real_sqrt_sum_squares_ge1 [simp]: "x \<le> sqrt (x\<^sup>2 + y\<^sup>2)"
by (rule power2_le_imp_le, simp_all)
-lemma real_sqrt_sum_squares_ge2 [simp]: "y \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma real_sqrt_sum_squares_ge2 [simp]: "y \<le> sqrt (x\<^sup>2 + y\<^sup>2)"
by (rule power2_le_imp_le, simp_all)
-lemma real_sqrt_ge_abs1 [simp]: "\<bar>x\<bar> \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma real_sqrt_ge_abs1 [simp]: "\<bar>x\<bar> \<le> sqrt (x\<^sup>2 + y\<^sup>2)"
by (rule power2_le_imp_le, simp_all)
-lemma real_sqrt_ge_abs2 [simp]: "\<bar>y\<bar> \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
+lemma real_sqrt_ge_abs2 [simp]: "\<bar>y\<bar> \<le> sqrt (x\<^sup>2 + y\<^sup>2)"
by (rule power2_le_imp_le, simp_all)
lemma le_real_sqrt_sumsq [simp]: "x \<le> sqrt (x * x + y * y)"
by (simp add: power2_eq_square [symmetric])
lemma real_sqrt_sum_squares_triangle_ineq:
- "sqrt ((a + c)\<twosuperior> + (b + d)\<twosuperior>) \<le> sqrt (a\<twosuperior> + b\<twosuperior>) + sqrt (c\<twosuperior> + d\<twosuperior>)"
+ "sqrt ((a + c)\<^sup>2 + (b + d)\<^sup>2) \<le> sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2)"
apply (rule power2_le_imp_le, simp)
apply (simp add: power2_sum)
apply (simp only: mult_assoc distrib_left [symmetric])
@@ -549,8 +549,8 @@
apply (rule power2_le_imp_le)
apply (simp add: power2_sum power_mult_distrib)
apply (simp add: ring_distribs)
-apply (subgoal_tac "0 \<le> b\<twosuperior> * c\<twosuperior> + a\<twosuperior> * d\<twosuperior> - 2 * (a * c) * (b * d)", simp)
-apply (rule_tac b="(a * d - b * c)\<twosuperior>" in ord_le_eq_trans)
+apply (subgoal_tac "0 \<le> b\<^sup>2 * c\<^sup>2 + a\<^sup>2 * d\<^sup>2 - 2 * (a * c) * (b * d)", simp)
+apply (rule_tac b="(a * d - b * c)\<^sup>2" in ord_le_eq_trans)
apply (rule zero_le_power2)
apply (simp add: power2_diff power_mult_distrib)
apply (simp add: mult_nonneg_nonneg)
@@ -559,7 +559,7 @@
done
lemma real_sqrt_sum_squares_less:
- "\<lbrakk>\<bar>x\<bar> < u / sqrt 2; \<bar>y\<bar> < u / sqrt 2\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) < u"
+ "\<lbrakk>\<bar>x\<bar> < u / sqrt 2; \<bar>y\<bar> < u / sqrt 2\<rbrakk> \<Longrightarrow> sqrt (x\<^sup>2 + y\<^sup>2) < u"
apply (rule power2_less_imp_less, simp)
apply (drule power_strict_mono [OF _ abs_ge_zero pos2])
apply (drule power_strict_mono [OF _ abs_ge_zero pos2])
@@ -571,12 +571,12 @@
text{*Needed for the infinitely close relation over the nonstandard
complex numbers*}
lemma lemma_sqrt_hcomplex_capprox:
- "[| 0 < u; x < u/2; y < u/2; 0 \<le> x; 0 \<le> y |] ==> sqrt (x\<twosuperior> + y\<twosuperior>) < u"
+ "[| 0 < u; x < u/2; y < u/2; 0 \<le> x; 0 \<le> y |] ==> sqrt (x\<^sup>2 + y\<^sup>2) < u"
apply (rule_tac y = "u/sqrt 2" in order_le_less_trans)
apply (erule_tac [2] lemma_real_divide_sqrt_less)
apply (rule power2_le_imp_le)
apply (auto simp add: zero_le_divide_iff power_divide)
-apply (rule_tac t = "u\<twosuperior>" in real_sum_of_halves [THEN subst])
+apply (rule_tac t = "u\<^sup>2" in real_sum_of_halves [THEN subst])
apply (rule add_mono)
apply (auto simp add: four_x_squared intro: power_mono)
done
--- a/src/HOL/Old_Number_Theory/Pocklington.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Old_Number_Theory/Pocklington.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1079,7 +1079,7 @@
shows "prime n"
unfolding prime_prime_factor_sqrt[of n]
proof-
- let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<twosuperior> \<le> n)"
+ let ?ths = "n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p\<^sup>2 \<le> n)"
from n have n01: "n\<noteq>0" "n\<noteq>1" by arith+
{fix p assume p: "prime p" "p dvd n" "p^2 \<le> n"
from p(3) sqr have "p^(Suc 1) \<le> q^(Suc 1)" by (simp add: power2_eq_square)
--- a/src/HOL/Old_Number_Theory/Primes.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Old_Number_Theory/Primes.thy Tue Aug 13 16:25:47 2013 +0200
@@ -39,7 +39,7 @@
lemma prime_dvd_square: "prime p ==> p dvd m^Suc (Suc 0) ==> p dvd m"
by (auto dest: prime_dvd_mult)
-lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
+lemma prime_dvd_power_two: "prime p ==> p dvd m\<^sup>2 ==> p dvd m"
by (rule prime_dvd_square) (simp_all add: power2_eq_square)
--- a/src/HOL/Power.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Power.thy Tue Aug 13 16:25:47 2013 +0200
@@ -27,14 +27,14 @@
text {* Special syntax for squares. *}
abbreviation (xsymbols)
- power2 :: "'a \<Rightarrow> 'a" ("(_\<twosuperior>)" [1000] 999) where
- "x\<twosuperior> \<equiv> x ^ 2"
+ power2 :: "'a \<Rightarrow> 'a" ("(_\<^sup>2)" [1000] 999) where
+ "x\<^sup>2 \<equiv> x ^ 2"
notation (latex output)
- power2 ("(_\<twosuperior>)" [1000] 999)
+ power2 ("(_\<^sup>2)" [1000] 999)
notation (HTML output)
- power2 ("(_\<twosuperior>)" [1000] 999)
+ power2 ("(_\<^sup>2)" [1000] 999)
end
@@ -67,7 +67,7 @@
"a ^ (m * n) = (a ^ m) ^ n"
by (induct n) (simp_all add: power_add)
-lemma power2_eq_square: "a\<twosuperior> = a * a"
+lemma power2_eq_square: "a\<^sup>2 = a * a"
by (simp add: numeral_2_eq_2)
lemma power3_eq_cube: "a ^ 3 = a * a * a"
@@ -139,10 +139,10 @@
lemma power_zero_numeral [simp]: "(0::'a) ^ numeral k = 0"
by (simp add: numeral_eq_Suc)
-lemma zero_power2: "0\<twosuperior> = 0" (* delete? *)
+lemma zero_power2: "0\<^sup>2 = 0" (* delete? *)
by (rule power_zero_numeral)
-lemma one_power2: "1\<twosuperior> = 1" (* delete? *)
+lemma one_power2: "1\<^sup>2 = 1" (* delete? *)
by (rule power_one)
end
@@ -218,7 +218,7 @@
by (simp only: neg_numeral_def power_minus_Bit1 power_numeral pow.simps)
lemma power2_minus [simp]:
- "(- a)\<twosuperior> = a\<twosuperior>"
+ "(- a)\<^sup>2 = a\<^sup>2"
by (rule power_minus_Bit0)
lemma power_minus1_even [simp]:
@@ -247,11 +247,11 @@
by (induct n) auto
lemma zero_eq_power2 [simp]:
- "a\<twosuperior> = 0 \<longleftrightarrow> a = 0"
+ "a\<^sup>2 = 0 \<longleftrightarrow> a = 0"
unfolding power2_eq_square by simp
lemma power2_eq_1_iff:
- "a\<twosuperior> = 1 \<longleftrightarrow> a = 1 \<or> a = - 1"
+ "a\<^sup>2 = 1 \<longleftrightarrow> a = 1 \<or> a = - 1"
unfolding power2_eq_square by (rule square_eq_1_iff)
end
@@ -259,7 +259,7 @@
context idom
begin
-lemma power2_eq_iff: "x\<twosuperior> = y\<twosuperior> \<longleftrightarrow> x = y \<or> x = - y"
+lemma power2_eq_iff: "x\<^sup>2 = y\<^sup>2 \<longleftrightarrow> x = y \<or> x = - y"
unfolding power2_eq_square by (rule square_eq_iff)
end
@@ -489,15 +489,15 @@
by (cases n) (simp_all del: power_Suc, rule power_inject_base)
lemma power2_le_imp_le:
- "x\<twosuperior> \<le> y\<twosuperior> \<Longrightarrow> 0 \<le> y \<Longrightarrow> x \<le> y"
+ "x\<^sup>2 \<le> y\<^sup>2 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x \<le> y"
unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)
lemma power2_less_imp_less:
- "x\<twosuperior> < y\<twosuperior> \<Longrightarrow> 0 \<le> y \<Longrightarrow> x < y"
+ "x\<^sup>2 < y\<^sup>2 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x < y"
by (rule power_less_imp_less_base)
lemma power2_eq_imp_eq:
- "x\<twosuperior> = y\<twosuperior> \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x = y"
+ "x\<^sup>2 = y\<^sup>2 \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x = y"
unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp
end
@@ -543,23 +543,23 @@
by (rule zero_le_power [OF abs_ge_zero])
lemma zero_le_power2 [simp]:
- "0 \<le> a\<twosuperior>"
+ "0 \<le> a\<^sup>2"
by (simp add: power2_eq_square)
lemma zero_less_power2 [simp]:
- "0 < a\<twosuperior> \<longleftrightarrow> a \<noteq> 0"
+ "0 < a\<^sup>2 \<longleftrightarrow> a \<noteq> 0"
by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)
lemma power2_less_0 [simp]:
- "\<not> a\<twosuperior> < 0"
+ "\<not> a\<^sup>2 < 0"
by (force simp add: power2_eq_square mult_less_0_iff)
lemma abs_power2 [simp]:
- "abs (a\<twosuperior>) = a\<twosuperior>"
+ "abs (a\<^sup>2) = a\<^sup>2"
by (simp add: power2_eq_square abs_mult abs_mult_self)
lemma power2_abs [simp]:
- "(abs a)\<twosuperior> = a\<twosuperior>"
+ "(abs a)\<^sup>2 = a\<^sup>2"
by (simp add: power2_eq_square abs_mult_self)
lemma odd_power_less_zero:
@@ -594,23 +594,23 @@
qed
lemma sum_power2_ge_zero:
- "0 \<le> x\<twosuperior> + y\<twosuperior>"
+ "0 \<le> x\<^sup>2 + y\<^sup>2"
by (intro add_nonneg_nonneg zero_le_power2)
lemma not_sum_power2_lt_zero:
- "\<not> x\<twosuperior> + y\<twosuperior> < 0"
+ "\<not> x\<^sup>2 + y\<^sup>2 < 0"
unfolding not_less by (rule sum_power2_ge_zero)
lemma sum_power2_eq_zero_iff:
- "x\<twosuperior> + y\<twosuperior> = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+ "x\<^sup>2 + y\<^sup>2 = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff)
lemma sum_power2_le_zero_iff:
- "x\<twosuperior> + y\<twosuperior> \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+ "x\<^sup>2 + y\<^sup>2 \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"
by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero)
lemma sum_power2_gt_zero_iff:
- "0 < x\<twosuperior> + y\<twosuperior> \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"
+ "0 < x\<^sup>2 + y\<^sup>2 \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"
unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff)
end
@@ -623,12 +623,12 @@
lemma power2_sum:
fixes x y :: "'a::comm_semiring_1"
- shows "(x + y)\<twosuperior> = x\<twosuperior> + y\<twosuperior> + 2 * x * y"
+ shows "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y"
by (simp add: algebra_simps power2_eq_square mult_2_right)
lemma power2_diff:
fixes x y :: "'a::comm_ring_1"
- shows "(x - y)\<twosuperior> = x\<twosuperior> + y\<twosuperior> - 2 * x * y"
+ shows "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y"
by (simp add: ring_distribs power2_eq_square mult_2) (rule mult_commute)
lemma power_0_Suc [simp]:
@@ -723,12 +723,12 @@
lemma power2_nat_le_eq_le:
fixes m n :: nat
- shows "m\<twosuperior> \<le> n\<twosuperior> \<longleftrightarrow> m \<le> n"
+ shows "m\<^sup>2 \<le> n\<^sup>2 \<longleftrightarrow> m \<le> n"
by (auto intro: power2_le_imp_le power_mono)
lemma power2_nat_le_imp_le:
fixes m n :: nat
- assumes "m\<twosuperior> \<le> n"
+ assumes "m\<^sup>2 \<le> n"
shows "m \<le> n"
using assms by (cases m) (simp_all add: power2_eq_square)
--- a/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Tue Aug 13 16:25:47 2013 +0200
@@ -26,16 +26,16 @@
datatype alphabet = a | b
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where
+ "[] \<in> S\<^sub>1"
+| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"
lemma
- "S\<^isub>1p w \<Longrightarrow> w = []"
+ "S\<^sub>1p w \<Longrightarrow> w = []"
quickcheck[tester = prolog, iterations=1, expect = counterexample]
oops
@@ -68,19 +68,19 @@
manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
-theorem S\<^isub>1_sound:
-"S\<^isub>1p w \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>1_sound:
+"S\<^sub>1p w \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = prolog, iterations=1, expect = counterexample]
oops
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where
+ "[] \<in> S\<^sub>2"
+| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2"
+| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2"
+| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
setup {* Code_Prolog.map_code_options (K
@@ -92,18 +92,18 @@
manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
-theorem S\<^isub>2_sound:
- "S\<^isub>2p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>2_sound:
+ "S\<^sub>2p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = prolog, iterations=1, expect = counterexample]
oops
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where
+ "[] \<in> S\<^sub>3"
+| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3"
+| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3"
+| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
setup {* Code_Prolog.map_code_options (K
@@ -114,8 +114,8 @@
replacing = [(("s3p", "limited_s3p"), "quickcheck")],
manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
-lemma S\<^isub>3_sound:
- "S\<^isub>3p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+lemma S\<^sub>3_sound:
+ "S\<^sub>3p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = prolog, iterations=1, size=1, expect = no_counterexample]
oops
@@ -132,20 +132,20 @@
prolog_system = Code_Prolog.SWI_PROLOG}) *}
-theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>3"
+theorem S\<^sub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>3"
quickcheck[tester = prolog, size=1, iterations=1]
oops
*)
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where
+ "[] \<in> S\<^sub>4"
+| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4"
+| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4"
+| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4"
+| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
setup {* Code_Prolog.map_code_options (K
@@ -157,14 +157,14 @@
manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
-theorem S\<^isub>4_sound:
- "S\<^isub>4p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>4_sound:
+ "S\<^sub>4p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = prolog, size=1, iterations=1, expect = no_counterexample]
oops
(*
-theorem S\<^isub>4_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+theorem S\<^sub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>4"
oops
*)
--- a/src/HOL/Predicate_Compile_Examples/Examples.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Examples.thy Tue Aug 13 16:25:47 2013 +0200
@@ -77,59 +77,59 @@
datatype alphabet = a | b
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where
+ "[] \<in> S\<^sub>1"
+| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"
-code_pred [inductify] S\<^isub>1p .
-code_pred [random_dseq inductify] S\<^isub>1p .
-thm S\<^isub>1p.equation
-thm S\<^isub>1p.random_dseq_equation
+code_pred [inductify] S\<^sub>1p .
+code_pred [random_dseq inductify] S\<^sub>1p .
+thm S\<^sub>1p.equation
+thm S\<^sub>1p.random_dseq_equation
-values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
+values [random_dseq 5, 5, 5] 5 "{x. S\<^sub>1p x}"
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where
+ "[] \<in> S\<^sub>2"
+| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2"
+| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2"
+| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
-code_pred [random_dseq inductify] S\<^isub>2p .
-thm S\<^isub>2p.random_dseq_equation
-thm A\<^isub>2p.random_dseq_equation
-thm B\<^isub>2p.random_dseq_equation
+code_pred [random_dseq inductify] S\<^sub>2p .
+thm S\<^sub>2p.random_dseq_equation
+thm A\<^sub>2p.random_dseq_equation
+thm B\<^sub>2p.random_dseq_equation
-values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
+values [random_dseq 5, 5, 5] 10 "{x. S\<^sub>2p x}"
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where
+ "[] \<in> S\<^sub>3"
+| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3"
+| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3"
+| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
-code_pred [inductify, skip_proof] S\<^isub>3p .
-thm S\<^isub>3p.equation
+code_pred [inductify, skip_proof] S\<^sub>3p .
+thm S\<^sub>3p.equation
-values 10 "{x. S\<^isub>3p x}"
+values 10 "{x. S\<^sub>3p x}"
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where
+ "[] \<in> S\<^sub>4"
+| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4"
+| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4"
+| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4"
+| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
-code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
+code_pred (expected_modes: o => bool, i => bool) S\<^sub>4p .
hide_const a b
@@ -272,7 +272,7 @@
fun aval :: "aexp \<Rightarrow> state' \<Rightarrow> nat" where
"aval (N n) _ = n" |
"aval (V x) st = st x" |
-"aval (Plus e\<^isub>1 e\<^isub>2) st = aval e\<^isub>1 st + aval e\<^isub>2 st"
+"aval (Plus e\<^sub>1 e\<^sub>2) st = aval e\<^sub>1 st + aval e\<^sub>2 st"
datatype bexp = B bool | Not bexp | And bexp bexp | Less aexp aexp
@@ -280,7 +280,7 @@
"bval (B b) _ = b" |
"bval (Not b) st = (\<not> bval b st)" |
"bval (And b1 b2) st = (bval b1 st \<and> bval b2 st)" |
-"bval (Less a\<^isub>1 a\<^isub>2) st = (aval a\<^isub>1 st < aval a\<^isub>2 st)"
+"bval (Less a\<^sub>1 a\<^sub>2) st = (aval a\<^sub>1 st < aval a\<^sub>2 st)"
datatype
com' = SKIP
@@ -295,14 +295,14 @@
Skip: "(SKIP,s) \<Rightarrow> s"
| Assign: "(x ::= a,s) \<Rightarrow> s(x := aval a s)"
-| Semi: "(c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2 \<Longrightarrow> (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3 \<Longrightarrow> (c\<^isub>1;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3"
+| Semi: "(c\<^sub>1,s\<^sub>1) \<Rightarrow> s\<^sub>2 \<Longrightarrow> (c\<^sub>2,s\<^sub>2) \<Rightarrow> s\<^sub>3 \<Longrightarrow> (c\<^sub>1;c\<^sub>2, s\<^sub>1) \<Rightarrow> s\<^sub>3"
-| IfTrue: "bval b s \<Longrightarrow> (c\<^isub>1,s) \<Rightarrow> t \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
-| IfFalse: "\<not>bval b s \<Longrightarrow> (c\<^isub>2,s) \<Rightarrow> t \<Longrightarrow> (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
+| IfTrue: "bval b s \<Longrightarrow> (c\<^sub>1,s) \<Rightarrow> t \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t"
+| IfFalse: "\<not>bval b s \<Longrightarrow> (c\<^sub>2,s) \<Rightarrow> t \<Longrightarrow> (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \<Rightarrow> t"
| WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s) \<Rightarrow> s"
-| WhileTrue: "bval b s\<^isub>1 \<Longrightarrow> (c,s\<^isub>1) \<Rightarrow> s\<^isub>2 \<Longrightarrow> (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3
- \<Longrightarrow> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3"
+| WhileTrue: "bval b s\<^sub>1 \<Longrightarrow> (c,s\<^sub>1) \<Rightarrow> s\<^sub>2 \<Longrightarrow> (WHILE b DO c, s\<^sub>2) \<Rightarrow> s\<^sub>3
+ \<Longrightarrow> (WHILE b DO c, s\<^sub>1) \<Rightarrow> s\<^sub>3"
code_pred big_step .
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Tue Aug 13 16:25:47 2013 +0200
@@ -76,9 +76,9 @@
definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
where
- "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
- s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
- no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
+ "feels_safe s r = (\<exists>s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'.
+ s = s\<^sub>3 @ [Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 \<and>
+ no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})"
section {* Some setup *}
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Tue Aug 13 16:25:47 2013 +0200
@@ -63,12 +63,12 @@
unfolding no_Check_in_def member_def by auto
lemma [code_pred_def]: "feels_safe s r =
-(EX s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
+(EX s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'.
s =
- s\<^isub>3 @
- [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 &
- no_Check_in (s\<^isub>3 @ s\<^isub>2) r &
- (\<not> (\<exists> g'. isinp (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r g')))"
+ s\<^sub>3 @
+ [Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 &
+ no_Check_in (s\<^sub>3 @ s\<^sub>2) r &
+ (\<not> (\<exists> g'. isinp (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r g')))"
unfolding feels_safe_def isinp_def by auto
setup {* Code_Prolog.map_code_options (K
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Tue Aug 13 16:25:47 2013 +0200
@@ -55,41 +55,41 @@
datatype alphabet = a | b
-inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
- "[] \<in> S\<^isub>1"
-| "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
-| "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
-| "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
+inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where
+ "[] \<in> S\<^sub>1"
+| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1"
+| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"
+| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"
lemma
- "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+ "w \<in> S\<^sub>1 \<Longrightarrow> w = []"
quickcheck[tester = predicate_compile_ff_nofs, iterations=1]
oops
-theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>1_sound:
+"w \<in> S\<^sub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator=predicate_compile_ff_nofs, size=15]
oops
-inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
- "[] \<in> S\<^isub>2"
-| "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
-| "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
-| "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
-| "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
+inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where
+ "[] \<in> S\<^sub>2"
+| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2"
+| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2"
+| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2"
+| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
(*
-code_pred [random_dseq inductify] S\<^isub>2 .
-thm S\<^isub>2.random_dseq_equation
-thm A\<^isub>2.random_dseq_equation
-thm B\<^isub>2.random_dseq_equation
+code_pred [random_dseq inductify] S\<^sub>2 .
+thm S\<^sub>2.random_dseq_equation
+thm A\<^sub>2.random_dseq_equation
+thm B\<^sub>2.random_dseq_equation
-values [random_dseq 1, 2, 8] 10 "{x. S\<^isub>2 x}"
+values [random_dseq 1, 2, 8] 10 "{x. S\<^sub>2 x}"
-lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
+lemma "w \<in> S\<^sub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
quickcheck[generator=predicate_compile, size=8]
oops
@@ -106,32 +106,32 @@
*)
lemma
-"w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
+"w \<in> S\<^sub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
quickcheck[generator=predicate_compile, size = 10, iterations = 1]
oops
*)
-theorem S\<^isub>2_sound:
-"w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>2_sound:
+"w \<in> S\<^sub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
oops
-inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
- "[] \<in> S\<^isub>3"
-| "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
-| "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
-| "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
-| "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
+inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where
+ "[] \<in> S\<^sub>3"
+| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3"
+| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3"
+| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3"
+| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
-code_pred [inductify, skip_proof] S\<^isub>3 .
-thm S\<^isub>3.equation
+code_pred [inductify, skip_proof] S\<^sub>3 .
+thm S\<^sub>3.equation
(*
-values 10 "{x. S\<^isub>3 x}"
+values 10 "{x. S\<^sub>3 x}"
*)
-lemma S\<^isub>3_sound:
-"w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+lemma S\<^sub>3_sound:
+"w \<in> S\<^sub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
oops
@@ -139,29 +139,29 @@
quickcheck[size=10, tester = predicate_compile_ff_fs]
oops
-theorem S\<^isub>3_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
+theorem S\<^sub>3_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^sub>3"
(*quickcheck[generator=SML]*)
quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
oops
-inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
- "[] \<in> S\<^isub>4"
-| "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
-| "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
-| "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
-| "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
-| "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
+inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where
+ "[] \<in> S\<^sub>4"
+| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4"
+| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4"
+| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4"
+| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4"
+| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
-theorem S\<^isub>4_sound:
-"w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+theorem S\<^sub>4_sound:
+"w \<in> S\<^sub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
quickcheck[tester = predicate_compile_ff_nofs, size=5, iterations=1]
oops
-theorem S\<^isub>4_complete:
-"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
+theorem S\<^sub>4_complete:
+"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^sub>4"
quickcheck[tester = predicate_compile_ff_nofs, size=5, iterations=1]
oops
--- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Tue Aug 13 16:25:47 2013 +0200
@@ -93,15 +93,15 @@
"[]\<langle>i\<rangle> = \<bottom>"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> \<lfloor>x\<rfloor> | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" ("_\<langle>_\<rangle>\<^isub>?" [90, 0] 91)
+primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" ("_\<langle>_\<rangle>\<^sub>?" [90, 0] 91)
where
- "[]\<langle>a\<rangle>\<^isub>? = \<bottom>"
-| "(x # xs)\<langle>a\<rangle>\<^isub>? = (if fst x = a then \<lfloor>snd x\<rfloor> else xs\<langle>a\<rangle>\<^isub>?)"
+ "[]\<langle>a\<rangle>\<^sub>? = \<bottom>"
+| "(x # xs)\<langle>a\<rangle>\<^sub>? = (if fst x = a then \<lfloor>snd x\<rfloor> else xs\<langle>a\<rangle>\<^sub>?)"
primrec unique :: "('a \<times> 'b) list \<Rightarrow> bool"
where
"unique [] = True"
-| "unique (x # xs) = (xs\<langle>fst x\<rangle>\<^isub>? = \<bottom> \<and> unique xs)"
+| "unique (x # xs) = (xs\<langle>fst x\<rangle>\<^sub>? = \<bottom> \<and> unique xs)"
datatype type =
TVar nat
@@ -132,66 +132,66 @@
| Abs type trm ("(3\<lambda>:_./ _)" [0, 10] 10)
| TAbs type trm ("(3\<lambda><:_./ _)" [0, 10] 10)
| App trm trm (infixl "\<bullet>" 200)
- | TApp trm type (infixl "\<bullet>\<^isub>\<tau>" 200)
+ | TApp trm type (infixl "\<bullet>\<^sub>\<tau>" 200)
-primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<up>\<^isub>\<tau>")
+primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<up>\<^sub>\<tau>")
where
- "\<up>\<^isub>\<tau> n k (TVar i) = (if i < k then TVar i else TVar (i + n))"
-| "\<up>\<^isub>\<tau> n k Top = Top"
-| "\<up>\<^isub>\<tau> n k (T \<rightarrow> U) = \<up>\<^isub>\<tau> n k T \<rightarrow> \<up>\<^isub>\<tau> n k U"
-| "\<up>\<^isub>\<tau> n k (\<forall><:T. U) = (\<forall><:\<up>\<^isub>\<tau> n k T. \<up>\<^isub>\<tau> n (k + 1) U)"
+ "\<up>\<^sub>\<tau> n k (TVar i) = (if i < k then TVar i else TVar (i + n))"
+| "\<up>\<^sub>\<tau> n k Top = Top"
+| "\<up>\<^sub>\<tau> n k (T \<rightarrow> U) = \<up>\<^sub>\<tau> n k T \<rightarrow> \<up>\<^sub>\<tau> n k U"
+| "\<up>\<^sub>\<tau> n k (\<forall><:T. U) = (\<forall><:\<up>\<^sub>\<tau> n k T. \<up>\<^sub>\<tau> n (k + 1) U)"
primrec lift :: "nat \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("\<up>")
where
"\<up> n k (Var i) = (if i < k then Var i else Var (i + n))"
-| "\<up> n k (\<lambda>:T. t) = (\<lambda>:\<up>\<^isub>\<tau> n k T. \<up> n (k + 1) t)"
-| "\<up> n k (\<lambda><:T. t) = (\<lambda><:\<up>\<^isub>\<tau> n k T. \<up> n (k + 1) t)"
+| "\<up> n k (\<lambda>:T. t) = (\<lambda>:\<up>\<^sub>\<tau> n k T. \<up> n (k + 1) t)"
+| "\<up> n k (\<lambda><:T. t) = (\<lambda><:\<up>\<^sub>\<tau> n k T. \<up> n (k + 1) t)"
| "\<up> n k (s \<bullet> t) = \<up> n k s \<bullet> \<up> n k t"
-| "\<up> n k (t \<bullet>\<^isub>\<tau> T) = \<up> n k t \<bullet>\<^isub>\<tau> \<up>\<^isub>\<tau> n k T"
+| "\<up> n k (t \<bullet>\<^sub>\<tau> T) = \<up> n k t \<bullet>\<^sub>\<tau> \<up>\<^sub>\<tau> n k T"
-primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("_[_ \<mapsto>\<^isub>\<tau> _]\<^isub>\<tau>" [300, 0, 0] 300)
+primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>\<tau>" [300, 0, 0] 300)
where
- "(TVar i)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> =
- (if k < i then TVar (i - 1) else if i = k then \<up>\<^isub>\<tau> k 0 S else TVar i)"
-| "Top[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = Top"
-| "(T \<rightarrow> U)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> \<rightarrow> U[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>"
-| "(\<forall><:T. U)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = (\<forall><:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. U[k+1 \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>)"
+ "(TVar i)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> =
+ (if k < i then TVar (i - 1) else if i = k then \<up>\<^sub>\<tau> k 0 S else TVar i)"
+| "Top[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = Top"
+| "(T \<rightarrow> U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> \<rightarrow> U[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>"
+| "(\<forall><:T. U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = (\<forall><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. U[k+1 \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>)"
-primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<down>\<^isub>\<tau>")
+primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<down>\<^sub>\<tau>")
where
- "\<down>\<^isub>\<tau> 0 k T = T"
-| "\<down>\<^isub>\<tau> (Suc n) k T = \<down>\<^isub>\<tau> n k (T[k \<mapsto>\<^isub>\<tau> Top]\<^isub>\<tau>)"
+ "\<down>\<^sub>\<tau> 0 k T = T"
+| "\<down>\<^sub>\<tau> (Suc n) k T = \<down>\<^sub>\<tau> n k (T[k \<mapsto>\<^sub>\<tau> Top]\<^sub>\<tau>)"
primrec subst :: "trm \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("_[_ \<mapsto> _]" [300, 0, 0] 300)
where
"(Var i)[k \<mapsto> s] = (if k < i then Var (i - 1) else if i = k then \<up> k 0 s else Var i)"
| "(t \<bullet> u)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet> u[k \<mapsto> s]"
-| "(t \<bullet>\<^isub>\<tau> T)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet>\<^isub>\<tau> \<down>\<^isub>\<tau> 1 k T"
-| "(\<lambda>:T. t)[k \<mapsto> s] = (\<lambda>:\<down>\<^isub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
-| "(\<lambda><:T. t)[k \<mapsto> s] = (\<lambda><:\<down>\<^isub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
+| "(t \<bullet>\<^sub>\<tau> T)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet>\<^sub>\<tau> \<down>\<^sub>\<tau> 1 k T"
+| "(\<lambda>:T. t)[k \<mapsto> s] = (\<lambda>:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
+| "(\<lambda><:T. t)[k \<mapsto> s] = (\<lambda><:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
-primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm" ("_[_ \<mapsto>\<^isub>\<tau> _]" [300, 0, 0] 300)
+primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm" ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
where
- "(Var i)[k \<mapsto>\<^isub>\<tau> S] = (if k < i then Var (i - 1) else Var i)"
-| "(t \<bullet> u)[k \<mapsto>\<^isub>\<tau> S] = t[k \<mapsto>\<^isub>\<tau> S] \<bullet> u[k \<mapsto>\<^isub>\<tau> S]"
-| "(t \<bullet>\<^isub>\<tau> T)[k \<mapsto>\<^isub>\<tau> S] = t[k \<mapsto>\<^isub>\<tau> S] \<bullet>\<^isub>\<tau> T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>"
-| "(\<lambda>:T. t)[k \<mapsto>\<^isub>\<tau> S] = (\<lambda>:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. t[k+1 \<mapsto>\<^isub>\<tau> S])"
-| "(\<lambda><:T. t)[k \<mapsto>\<^isub>\<tau> S] = (\<lambda><:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. t[k+1 \<mapsto>\<^isub>\<tau> S])"
+ "(Var i)[k \<mapsto>\<^sub>\<tau> S] = (if k < i then Var (i - 1) else Var i)"
+| "(t \<bullet> u)[k \<mapsto>\<^sub>\<tau> S] = t[k \<mapsto>\<^sub>\<tau> S] \<bullet> u[k \<mapsto>\<^sub>\<tau> S]"
+| "(t \<bullet>\<^sub>\<tau> T)[k \<mapsto>\<^sub>\<tau> S] = t[k \<mapsto>\<^sub>\<tau> S] \<bullet>\<^sub>\<tau> T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>"
+| "(\<lambda>:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda>:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"
+| "(\<lambda><:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"
-primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<up>\<^isub>e")
+primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<up>\<^sub>e")
where
- "\<up>\<^isub>e n k [] = []"
-| "\<up>\<^isub>e n k (B \<Colon> \<Gamma>) = mapB (\<up>\<^isub>\<tau> n (k + \<parallel>\<Gamma>\<parallel>)) B \<Colon> \<up>\<^isub>e n k \<Gamma>"
+ "\<up>\<^sub>e n k [] = []"
+| "\<up>\<^sub>e n k (B \<Colon> \<Gamma>) = mapB (\<up>\<^sub>\<tau> n (k + \<parallel>\<Gamma>\<parallel>)) B \<Colon> \<up>\<^sub>e n k \<Gamma>"
-primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env" ("_[_ \<mapsto>\<^isub>\<tau> _]\<^isub>e" [300, 0, 0] 300)
+primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env" ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>e" [300, 0, 0] 300)
where
- "[][k \<mapsto>\<^isub>\<tau> T]\<^isub>e = []"
-| "(B \<Colon> \<Gamma>)[k \<mapsto>\<^isub>\<tau> T]\<^isub>e = mapB (\<lambda>U. U[k + \<parallel>\<Gamma>\<parallel> \<mapsto>\<^isub>\<tau> T]\<^isub>\<tau>) B \<Colon> \<Gamma>[k \<mapsto>\<^isub>\<tau> T]\<^isub>e"
+ "[][k \<mapsto>\<^sub>\<tau> T]\<^sub>e = []"
+| "(B \<Colon> \<Gamma>)[k \<mapsto>\<^sub>\<tau> T]\<^sub>e = mapB (\<lambda>U. U[k + \<parallel>\<Gamma>\<parallel> \<mapsto>\<^sub>\<tau> T]\<^sub>\<tau>) B \<Colon> \<Gamma>[k \<mapsto>\<^sub>\<tau> T]\<^sub>e"
-primrec decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<down>\<^isub>e")
+primrec decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<down>\<^sub>e")
where
- "\<down>\<^isub>e 0 k \<Gamma> = \<Gamma>"
-| "\<down>\<^isub>e (Suc n) k \<Gamma> = \<down>\<^isub>e n k (\<Gamma>[k \<mapsto>\<^isub>\<tau> Top]\<^isub>e)"
+ "\<down>\<^sub>e 0 k \<Gamma> = \<Gamma>"
+| "\<down>\<^sub>e (Suc n) k \<Gamma> = \<down>\<^sub>e n k (\<Gamma>[k \<mapsto>\<^sub>\<tau> Top]\<^sub>e)"
inductive
well_formed :: "env \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile>\<^bsub>wf\<^esub> _" [50, 50] 50)
@@ -224,20 +224,20 @@
SA_Top: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> S \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
| SA_refl_TVar: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> TVar i \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: TVar i"
| SA_trans_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB U\<rfloor> \<Longrightarrow>
- \<Gamma> \<turnstile> \<up>\<^isub>\<tau> (Suc i) 0 U <: T \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: T"
-| SA_arrow: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2"
-| SA_all: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<Longrightarrow> TVarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2 \<Longrightarrow>
- \<Gamma> \<turnstile> (\<forall><:S\<^isub>1. S\<^isub>2) <: (\<forall><:T\<^isub>1. T\<^isub>2)"
+ \<Gamma> \<turnstile> \<up>\<^sub>\<tau> (Suc i) 0 U <: T \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: T"
+| SA_arrow: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<Longrightarrow> \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2"
+| SA_all: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<Longrightarrow> TVarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2 \<Longrightarrow>
+ \<Gamma> \<turnstile> (\<forall><:S\<^sub>1. S\<^sub>2) <: (\<forall><:T\<^sub>1. T\<^sub>2)"
inductive
typing :: "env \<Rightarrow> trm \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
where
- T_Var: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma>\<langle>i\<rangle> = \<lfloor>VarB U\<rfloor> \<Longrightarrow> T = \<up>\<^isub>\<tau> (Suc i) 0 U \<Longrightarrow> \<Gamma> \<turnstile> Var i : T"
-| T_Abs: "VarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \<rightarrow> \<down>\<^isub>\<tau> 1 0 T\<^isub>2"
-| T_App: "\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>11 \<rightarrow> T\<^isub>12 \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>11 \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<bullet> t\<^isub>2 : T\<^isub>12"
-| T_TAbs: "TVarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda><:T\<^isub>1. t\<^isub>2) : (\<forall><:T\<^isub>1. T\<^isub>2)"
-| T_TApp: "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall><:T\<^isub>11. T\<^isub>12) \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>11 \<Longrightarrow>
- \<Gamma> \<turnstile> t\<^isub>1 \<bullet>\<^isub>\<tau> T\<^isub>2 : T\<^isub>12[0 \<mapsto>\<^isub>\<tau> T\<^isub>2]\<^isub>\<tau>"
+ T_Var: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma>\<langle>i\<rangle> = \<lfloor>VarB U\<rfloor> \<Longrightarrow> T = \<up>\<^sub>\<tau> (Suc i) 0 U \<Longrightarrow> \<Gamma> \<turnstile> Var i : T"
+| T_Abs: "VarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \<rightarrow> \<down>\<^sub>\<tau> 1 0 T\<^sub>2"
+| T_App: "\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>11 \<rightarrow> T\<^sub>12 \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11 \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<bullet> t\<^sub>2 : T\<^sub>12"
+| T_TAbs: "TVarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda><:T\<^sub>1. t\<^sub>2) : (\<forall><:T\<^sub>1. T\<^sub>2)"
+| T_TApp: "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall><:T\<^sub>11. T\<^sub>12) \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11 \<Longrightarrow>
+ \<Gamma> \<turnstile> t\<^sub>1 \<bullet>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[0 \<mapsto>\<^sub>\<tau> T\<^sub>2]\<^sub>\<tau>"
| T_Sub: "\<Gamma> \<turnstile> t : S \<Longrightarrow> \<Gamma> \<turnstile> S <: T \<Longrightarrow> \<Gamma> \<turnstile> t : T"
code_pred [inductify, skip_proof, specialise] typing .
--- a/src/HOL/Probability/Binary_Product_Measure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -16,36 +16,36 @@
section "Binary products"
-definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
- "A \<Otimes>\<^isub>M B = measure_of (space A \<times> space B)
+definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
+ "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
- (\<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
+ (\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
using sets.space_closed[of A] sets.space_closed[of B] by auto
lemma space_pair_measure:
- "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
+ "space (A \<Otimes>\<^sub>M B) = space A \<times> space B"
unfolding pair_measure_def using pair_measure_closed[of A B]
by (rule space_measure_of)
lemma sets_pair_measure:
- "sets (A \<Otimes>\<^isub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+ "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
unfolding pair_measure_def using pair_measure_closed[of A B]
by (rule sets_measure_of)
lemma sets_pair_measure_cong[cong]:
- "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')"
+ "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"
unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
lemma pair_measureI[intro, simp, measurable]:
- "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
+ "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"
by (auto simp: sets_pair_measure)
lemma measurable_pair_measureI:
assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M"
- shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+ shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
unfolding pair_measure_def using 1 2
by (intro measurable_measure_of) (auto dest: sets.sets_into_space)
@@ -55,7 +55,7 @@
lemma measurable_Pair[measurable (raw)]:
assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2"
- shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+ shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
proof (rule measurable_pair_measureI)
show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"
using f g by (auto simp: measurable_def)
@@ -68,60 +68,60 @@
qed
lemma measurable_Pair_compose_split[measurable_dest]:
- assumes f: "split f \<in> measurable (M1 \<Otimes>\<^isub>M M2) N"
+ assumes f: "split f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"
assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"
shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"
using measurable_compose[OF measurable_Pair f, OF g h] by simp
lemma measurable_pair:
assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
- shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+ shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
using measurable_Pair[OF assms] by simp
-lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^isub>M M2) M1"
+lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
measurable_def)
-lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^isub>M M2) M2"
+lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
measurable_def)
lemma
- assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^isub>M P)"
+ assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"
shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"
and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"
by simp_all
lemma
assumes f[measurable]: "f \<in> measurable M N"
- shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^isub>M P) N"
- and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^isub>M M) N"
+ shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N"
+ and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"
by simp_all
lemma measurable_pair_iff:
- "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+ "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
by (auto intro: measurable_pair[of f M M1 M2])
lemma measurable_split_conv:
"(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"
by (intro arg_cong2[where f="op \<in>"]) auto
-lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
+lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
by (auto intro!: measurable_Pair simp: measurable_split_conv)
lemma measurable_pair_swap:
- assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M"
using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)
lemma measurable_pair_swap_iff:
- "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) M"
+ "f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M"
by (auto dest: measurable_pair_swap)
-lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^isub>M M2)"
+lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
by simp
lemma sets_Pair1[measurable (raw)]:
- assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "Pair x -` A \<in> sets M2"
+ assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"
proof -
have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
@@ -130,10 +130,10 @@
finally show ?thesis .
qed
-lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^isub>M M2)"
+lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
by (auto intro!: measurable_Pair)
-lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
+lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
proof -
have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
@@ -143,13 +143,13 @@
qed
lemma measurable_Pair2:
- assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and x: "x \<in> space M1"
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"
shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
using measurable_comp[OF measurable_Pair1' f, OF x]
by (simp add: comp_def)
lemma measurable_Pair1:
- assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and y: "y \<in> space M2"
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"
shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
using measurable_comp[OF measurable_Pair2' f, OF y]
by (simp add: comp_def)
@@ -162,7 +162,7 @@
by (auto simp: disjoint_family_on_def)
lemma (in finite_measure) finite_measure_cut_measurable:
- assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^isub>M M)"
+ assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"
(is "?s Q \<in> _")
using Int_stable_pair_measure_generator pair_measure_closed assms
@@ -184,7 +184,7 @@
qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)
lemma (in sigma_finite_measure) measurable_emeasure_Pair:
- assumes Q: "Q \<in> sets (N \<Otimes>\<^isub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
+ assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
proof -
from sigma_finite_disjoint guess F . note F = this
then have F_sets: "\<And>i. F i \<in> sets M" by auto
@@ -209,43 +209,43 @@
have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"
proof (intro suminf_emeasure)
show "range (?C x) \<subseteq> sets M"
- using F `Q \<in> sets (N \<Otimes>\<^isub>M M)` by (auto intro!: sets_Pair1)
+ using F `Q \<in> sets (N \<Otimes>\<^sub>M M)` by (auto intro!: sets_Pair1)
have "disjoint_family F" using F by auto
show "disjoint_family (?C x)"
by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
qed
also have "(\<Union>i. ?C x i) = Pair x -` Q"
- using F sets.sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^isub>M M)`]
+ using F sets.sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^sub>M M)`]
by (auto simp: space_pair_measure)
finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"
by simp }
- ultimately show ?thesis using `Q \<in> sets (N \<Otimes>\<^isub>M M)` F_sets
+ ultimately show ?thesis using `Q \<in> sets (N \<Otimes>\<^sub>M M)` F_sets
by auto
qed
lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M"
- assumes A: "{x\<in>space (N \<Otimes>\<^isub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^isub>M M)"
+ assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"
proof -
- from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^isub>M M). snd x \<in> A (fst x)} = A x"
+ from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x"
by (auto simp: space_pair_measure)
with measurable_emeasure_Pair[OF A] show ?thesis
by (auto cong: measurable_cong)
qed
lemma (in sigma_finite_measure) emeasure_pair_measure:
- assumes "X \<in> sets (N \<Otimes>\<^isub>M M)"
- shows "emeasure (N \<Otimes>\<^isub>M M) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
+ assumes "X \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
proof (rule emeasure_measure_of[OF pair_measure_def])
- show "positive (sets (N \<Otimes>\<^isub>M M)) ?\<mu>"
+ show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
by (auto simp: positive_def positive_integral_positive)
have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
by (auto simp: indicator_def)
- show "countably_additive (sets (N \<Otimes>\<^isub>M M)) ?\<mu>"
+ show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
proof (rule countably_additiveI)
- fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^isub>M M)" "disjoint_family F"
- from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^isub>M M)" "(\<Union>i. F i) \<in> sets (N \<Otimes>\<^isub>M M)" by auto
+ fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
+ from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" "(\<Union>i. F i) \<in> sets (N \<Otimes>\<^sub>M M)" by auto
moreover from F have "\<And>i. (\<lambda>x. emeasure M (Pair x -` F i)) \<in> borel_measurable N"
by (intro measurable_emeasure_Pair) auto
moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
@@ -261,8 +261,8 @@
qed fact
lemma (in sigma_finite_measure) emeasure_pair_measure_alt:
- assumes X: "X \<in> sets (N \<Otimes>\<^isub>M M)"
- shows "emeasure (N \<Otimes>\<^isub>M M) X = (\<integral>\<^isup>+x. emeasure M (Pair x -` X) \<partial>N)"
+ assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"
proof -
have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"
by (auto simp: indicator_def)
@@ -272,9 +272,9 @@
lemma (in sigma_finite_measure) emeasure_pair_measure_Times:
assumes A: "A \<in> sets N" and B: "B \<in> sets M"
- shows "emeasure (N \<Otimes>\<^isub>M M) (A \<times> B) = emeasure N A * emeasure M B"
+ shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"
proof -
- have "emeasure (N \<Otimes>\<^isub>M M) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M B * indicator A x \<partial>N)"
+ have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"
using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)
also have "\<dots> = emeasure M B * emeasure N A"
using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)
@@ -288,16 +288,16 @@
for M1 :: "'a measure" and M2 :: "'b measure"
lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
- "Q \<in> sets (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
+ "Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
using M2.measurable_emeasure_Pair .
lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
- assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+ assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
proof -
- have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+ have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
using Q measurable_pair_swap' by (auto intro: measurable_sets)
note M1.measurable_emeasure_Pair[OF this]
- moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1)) = (\<lambda>x. (x, y)) -` Q"
+ moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q"
using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
ultimately show ?thesis by simp
qed
@@ -305,7 +305,7 @@
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
- (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
+ (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
proof -
from M1.sigma_finite_incseq guess F1 . note F1 = this
from M2.sigma_finite_incseq guess F2 . note F2 = this
@@ -336,26 +336,26 @@
fix i
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]
- show "emeasure (M1 \<Otimes>\<^isub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
+ show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
by (auto simp add: emeasure_pair_measure_Times)
qed
qed
-sublocale pair_sigma_finite \<subseteq> P: sigma_finite_measure "M1 \<Otimes>\<^isub>M M2"
+sublocale pair_sigma_finite \<subseteq> P: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
proof
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
- show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
+ show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
proof (rule exI[of _ F], intro conjI)
- show "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" using F by (auto simp: pair_measure_def)
- show "(\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2)"
+ show "range F \<subseteq> sets (M1 \<Otimes>\<^sub>M M2)" using F by (auto simp: pair_measure_def)
+ show "(\<Union>i. F i) = space (M1 \<Otimes>\<^sub>M M2)"
using F by (auto simp: space_pair_measure)
- show "\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>" using F by auto
+ show "\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>" using F by auto
qed
qed
lemma sigma_finite_pair_measure:
assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
- shows "sigma_finite_measure (A \<Otimes>\<^isub>M B)"
+ shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"
proof -
interpret A: sigma_finite_measure A by fact
interpret B: sigma_finite_measure B by fact
@@ -364,12 +364,12 @@
qed
lemma sets_pair_swap:
- assumes "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
- shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+ assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
using measurable_pair_swap' assms by (rule measurable_sets)
lemma (in pair_sigma_finite) distr_pair_swap:
- "M1 \<Otimes>\<^isub>M M2 = distr (M2 \<Otimes>\<^isub>M M1) (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
+ "M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
proof -
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
@@ -387,20 +387,20 @@
next
fix X assume "X \<in> ?E"
then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
- have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^isub>M M1) = B \<times> A"
+ have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A"
using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure)
- with A B show "emeasure (M1 \<Otimes>\<^isub>M M2) X = emeasure ?D X"
+ with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X"
by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr
measurable_pair_swap' ac_simps)
qed
qed
lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
- assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
- shows "emeasure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
+ assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
(is "_ = ?\<nu> A")
proof -
- have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))) = (\<lambda>x. (x, y)) -` A"
+ have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A"
using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
show ?thesis using A
by (subst distr_pair_swap)
@@ -409,14 +409,14 @@
qed
lemma (in pair_sigma_finite) AE_pair:
- assumes "AE x in (M1 \<Otimes>\<^isub>M M2). Q x"
+ assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"
shows "AE x in M1. (AE y in M2. Q (x, y))"
proof -
- obtain N where N: "N \<in> sets (M1 \<Otimes>\<^isub>M M2)" "emeasure (M1 \<Otimes>\<^isub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> Q x} \<subseteq> N"
+ obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N"
using assms unfolding eventually_ae_filter by auto
show ?thesis
proof (rule AE_I)
- from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
+ from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^sub>M M2)`]
show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
by (auto simp: M2.emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)
show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
@@ -435,41 +435,41 @@
qed
lemma (in pair_sigma_finite) AE_pair_measure:
- assumes "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+ assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
assumes ae: "AE x in M1. AE y in M2. P (x, y)"
- shows "AE x in M1 \<Otimes>\<^isub>M M2. P x"
+ shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"
proof (subst AE_iff_measurable[OF _ refl])
- show "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+ show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
by (rule sets.sets_Collect) fact
- then have "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} =
- (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
+ then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
+ (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
by (simp add: M2.emeasure_pair_measure)
- also have "\<dots> = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. 0 \<partial>M2 \<partial>M1)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)"
using ae
apply (safe intro!: positive_integral_cong_AE)
apply (intro AE_I2)
apply (safe intro!: positive_integral_cong_AE)
apply auto
done
- finally show "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} = 0" by simp
+ finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
qed
lemma (in pair_sigma_finite) AE_pair_iff:
- "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow>
- (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x))"
+ "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
+ (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))"
using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
lemma (in pair_sigma_finite) AE_commute:
- assumes P: "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+ assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
proof -
interpret Q: pair_sigma_finite M2 M1 ..
have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x"
by auto
- have "{x \<in> space (M2 \<Otimes>\<^isub>M M1). P (snd x) (fst x)} =
- (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^isub>M M1)"
+ have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} =
+ (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)"
by (auto simp: space_pair_measure)
- also have "\<dots> \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+ also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)"
by (intro sets_pair_swap P)
finally show ?thesis
apply (subst AE_pair_iff[OF P])
@@ -483,12 +483,12 @@
section "Fubinis theorem"
lemma measurable_compose_Pair1:
- "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^isub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
+ "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
by simp
lemma (in sigma_finite_measure) borel_measurable_positive_integral_fst':
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M)" "\<And>x. 0 \<le> f x"
- shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
+ shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
using f proof induct
case (cong u v)
then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"
@@ -502,7 +502,7 @@
case (set Q)
have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y"
by (auto simp: indicator_def)
- have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^isup>+ y. indicator Q (x, y) \<partial>M"
+ have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M"
by (simp add: sets_Pair1[OF set])
from this measurable_emeasure_Pair[OF set] show ?case
by (rule measurable_cong[THEN iffD1])
@@ -511,8 +511,8 @@
cong: measurable_cong)
lemma (in sigma_finite_measure) positive_integral_fst:
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M)" "\<And>x. 0 \<le> f x"
- shows "(\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M) f" (is "?I f = _")
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"
+ shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")
using f proof induct
case (cong u v)
moreover then have "?I u = ?I v"
@@ -526,46 +526,46 @@
cong: positive_integral_cong)
lemma (in sigma_finite_measure) positive_integral_fst_measurable:
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M)"
- shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
+ shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
(is "?C f \<in> borel_measurable M1")
- and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M) f"
+ and "(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M) f"
using f
borel_measurable_positive_integral_fst'[of "\<lambda>x. max 0 (f x)"]
positive_integral_fst[of "\<lambda>x. max 0 (f x)"]
unfolding positive_integral_max_0 by auto
lemma (in sigma_finite_measure) borel_measurable_positive_integral[measurable (raw)]:
- "split f \<in> borel_measurable (N \<Otimes>\<^isub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M) \<in> borel_measurable N"
+ "split f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"
using positive_integral_fst_measurable(1)[of "split f" N] by simp
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
- "split f \<in> borel_measurable (N \<Otimes>\<^isub>M M) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M) \<in> borel_measurable N"
+ "split f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M) \<in> borel_measurable N"
by (simp add: lebesgue_integral_def)
lemma (in pair_sigma_finite) positive_integral_snd_measurable:
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
- shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) f"
proof -
note measurable_pair_swap[OF f]
from M1.positive_integral_fst_measurable[OF this]
- have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))"
+ have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"
by simp
- also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+ also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) f"
by (subst distr_pair_swap)
(auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)
finally show ?thesis .
qed
lemma (in pair_sigma_finite) Fubini:
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
- shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
unfolding positive_integral_snd_measurable[OF assms]
unfolding M2.positive_integral_fst_measurable[OF assms] ..
lemma (in pair_sigma_finite) integrable_product_swap:
- assumes "integrable (M1 \<Otimes>\<^isub>M M2) f"
- shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
+ assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"
+ shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"
proof -
interpret Q: pair_sigma_finite M2 M1 by default
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
@@ -575,7 +575,7 @@
qed
lemma (in pair_sigma_finite) integrable_product_swap_iff:
- "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^isub>M M2) f"
+ "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f"
proof -
interpret Q: pair_sigma_finite M2 M1 by default
from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
@@ -583,8 +583,8 @@
qed
lemma (in pair_sigma_finite) integral_product_swap:
- assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
- shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
proof -
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
show ?thesis unfolding *
@@ -592,67 +592,67 @@
qed
lemma (in pair_sigma_finite) integrable_fst_measurable:
- assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "AE x in M1. integrable M2 (\<lambda> y. f (x, y))" (is "?AE")
- and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
+ and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" (is "?INT")
proof -
- have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+ have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
using f by auto
let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)"
have
- borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" and
- int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?nf \<noteq> \<infinity>" "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?pf \<noteq> \<infinity>"
+ borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" and
+ int: "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) ?nf \<noteq> \<infinity>" "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) ?pf \<noteq> \<infinity>"
using assms by auto
- have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
- "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+ have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
+ "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
using borel[THEN M2.positive_integral_fst_measurable(1)] int
unfolding borel[THEN M2.positive_integral_fst_measurable(2)] by simp_all
with borel[THEN M2.positive_integral_fst_measurable(1)]
- have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
- "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+ have AE_pos: "AE x in M1. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+ "AE x in M1. (\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
by (auto intro!: positive_integral_PInf_AE )
- then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
- "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+ then have AE: "AE x in M1. \<bar>\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
+ "AE x in M1. \<bar>\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
by (auto simp: positive_integral_positive)
from AE_pos show ?AE using assms
by (simp add: measurable_Pair2[OF f_borel] integrable_def)
- { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+ { fix f have "(\<integral>\<^sup>+ x. - \<integral>\<^sup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^sup>+x. 0 \<partial>M1)"
using positive_integral_positive
by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
- then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
+ then have "(\<integral>\<^sup>+ x. - \<integral>\<^sup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
note this[simp]
- { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
- and int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
- and AE: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
- have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
+ { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ and int: "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
+ and AE: "AE x in M1. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+ have "integrable M1 (\<lambda>x. real (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
proof (intro integrable_def[THEN iffD2] conjI)
show "?f \<in> borel_measurable M1"
using borel by (auto intro!: M2.positive_integral_fst_measurable)
- have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1)"
+ have "(\<integral>\<^sup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1)"
using AE positive_integral_positive[of M2]
by (auto intro!: positive_integral_cong_AE simp: ereal_real)
- then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
+ then show "(\<integral>\<^sup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
using M2.positive_integral_fst_measurable[OF borel] int by simp
- have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
+ have "(\<integral>\<^sup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^sup>+x. 0 \<partial>M1)"
by (intro positive_integral_cong_pos)
(simp add: positive_integral_positive real_of_ereal_pos)
- then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
+ then show "(\<integral>\<^sup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
qed }
with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
show ?INT
- unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^isub>M M2"] lebesgue_integral_def[of M2]
+ unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^sub>M M2"] lebesgue_integral_def[of M2]
borel[THEN M2.positive_integral_fst_measurable(2), symmetric]
using AE[THEN integral_real]
by simp
qed
lemma (in pair_sigma_finite) integrable_snd_measurable:
- assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "AE y in M2. integrable M1 (\<lambda>x. f (x, y))" (is "?AE")
- and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
+ and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" (is "?INT")
proof -
interpret Q: pair_sigma_finite M2 M1 by default
- have Q_int: "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x, y). f (y, x))"
+ have Q_int: "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x, y). f (y, x))"
using f unfolding integrable_product_swap_iff .
show ?INT
using Q.integrable_fst_measurable(2)[OF Q_int]
@@ -663,7 +663,7 @@
qed
lemma (in pair_sigma_finite) Fubini_integral:
- assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
unfolding integrable_snd_measurable[OF assms]
unfolding integrable_fst_measurable[OF assms] ..
@@ -696,7 +696,7 @@
lemma pair_measure_count_space:
assumes A: "finite A" and B: "finite B"
- shows "count_space A \<Otimes>\<^isub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
+ shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
proof (rule measure_eqI)
interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
@@ -729,14 +729,14 @@
assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"
assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"
- shows "density M1 f \<Otimes>\<^isub>M density M2 g = density (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
+ shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
proof (rule measure_eqI)
interpret M2: sigma_finite_measure M2 by fact
interpret D2: sigma_finite_measure "density M2 g" by fact
fix A assume A: "A \<in> sets ?L"
- with f g have "(\<integral>\<^isup>+ x. f x * \<integral>\<^isup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) =
- (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"
+ with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) =
+ (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"
by (intro positive_integral_cong_AE)
(auto simp add: positive_integral_cmult[symmetric] ac_simps)
with A f g show "emeasure ?L A = emeasure ?R A"
@@ -765,7 +765,7 @@
lemma pair_measure_distr:
assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
assumes "sigma_finite_measure (distr N T g)"
- shows "distr M S f \<Otimes>\<^isub>M distr N T g = distr (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
+ shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
proof (rule measure_eqI)
interpret T: sigma_finite_measure "distr N T g" by fact
interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+
@@ -779,16 +779,16 @@
lemma pair_measure_eqI:
assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
- assumes sets: "sets (M1 \<Otimes>\<^isub>M M2) = sets M"
+ assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"
assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"
- shows "M1 \<Otimes>\<^isub>M M2 = M"
+ shows "M1 \<Otimes>\<^sub>M M2 = M"
proof -
interpret M1: sigma_finite_measure M1 by fact
interpret M2: sigma_finite_measure M2 by fact
interpret pair_sigma_finite M1 M2 by default
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
- let ?P = "M1 \<Otimes>\<^isub>M M2"
+ let ?P = "M1 \<Otimes>\<^sub>M M2"
show ?thesis
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
show "?E \<subseteq> Pow (space ?P)"
--- a/src/HOL/Probability/Distributions.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Distributions.thy Tue Aug 13 16:25:47 2013 +0200
@@ -54,7 +54,7 @@
have deriv: "\<And>x. DERIV ?F x :> ?f x" "\<And>x. 0 \<le> l * exp (- x * l)"
by (auto intro!: DERIV_intros simp: zero_le_mult_iff)
- have "emeasure ?D (space ?D) = (\<integral>\<^isup>+ x. ereal (?f x) * indicator {0..} x \<partial>lborel)"
+ have "emeasure ?D (space ?D) = (\<integral>\<^sup>+ x. ereal (?f x) * indicator {0..} x \<partial>lborel)"
by (auto simp: emeasure_density exponential_density_def
intro!: positive_integral_cong split: split_indicator)
also have "\<dots> = ereal (0 - ?F 0)"
@@ -71,10 +71,10 @@
then show "prob_space ?D" by rule
assume "0 \<le> a"
- have "emeasure ?D {..a} = (\<integral>\<^isup>+x. ereal (?f x) * indicator {0..a} x \<partial>lborel)"
+ have "emeasure ?D {..a} = (\<integral>\<^sup>+x. ereal (?f x) * indicator {0..a} x \<partial>lborel)"
by (auto simp add: emeasure_density intro!: positive_integral_cong split: split_indicator)
(auto simp: exponential_density_def)
- also have "(\<integral>\<^isup>+x. ereal (?f x) * indicator {0..a} x \<partial>lborel) = ereal (?F a) - ereal (?F 0)"
+ also have "(\<integral>\<^sup>+x. ereal (?f x) * indicator {0..a} x \<partial>lborel) = ereal (?F a) - ereal (?F 0)"
using `0 \<le> a` deriv by (intro positive_integral_FTC_atLeastAtMost) auto
also have "\<dots> = 1 - exp (- a * l)"
by simp
@@ -141,7 +141,7 @@
have "\<And>x. \<not> 0 \<le> a \<Longrightarrow> ereal (exponential_density l x) * indicator {..a} x = 0"
by (simp add: exponential_density_def)
- then show "(\<integral>\<^isup>+ x. exponential_density l x * indicator {..a} x \<partial>lborel) = ereal (if 0 \<le> a then 1 - exp (- a * l) else 0)"
+ then show "(\<integral>\<^sup>+ x. exponential_density l x * indicator {..a} x \<partial>lborel) = ereal (if 0 \<le> a then 1 - exp (- a * l) else 0)"
using emeasure_exponential_density_le0[of l a]
by (auto simp: emeasure_density times_ereal.simps[symmetric] ereal_indicator
simp del: times_ereal.simps ereal_zero_times)
@@ -167,10 +167,10 @@
let ?f = "\<lambda>i x. x * exp (- x) * indicator {0::real .. i} x"
have "eventually (\<lambda>b::real. 0 \<le> b) at_top"
by (rule eventually_ge_at_top)
- then have "eventually (\<lambda>b. 1 - (inverse (exp b) + b / exp b) = integral\<^isup>L lborel (?f b)) at_top"
+ then have "eventually (\<lambda>b. 1 - (inverse (exp b) + b / exp b) = integral\<^sup>L lborel (?f b)) at_top"
proof eventually_elim
fix b :: real assume [simp]: "0 \<le> b"
- have "(\<integral>x. (exp (-x)) * indicator {0 .. b} x \<partial>lborel) - (integral\<^isup>L lborel (?f b)) =
+ have "(\<integral>x. (exp (-x)) * indicator {0 .. b} x \<partial>lborel) - (integral\<^sup>L lborel (?f b)) =
(\<integral>x. (exp (-x) - x * exp (-x)) * indicator {0 .. b} x \<partial>lborel)"
by (subst integral_diff(2)[symmetric])
(auto intro!: borel_integrable_atLeastAtMost integral_cong split: split_indicator)
@@ -184,16 +184,16 @@
qed fact
also have "(\<integral>x. (exp (-x)) * indicator {0 .. b} x \<partial>lborel) = - exp (- b) - - exp (- 0)"
by (rule integral_FTC_atLeastAtMost) (auto intro!: DERIV_intros)
- finally show "1 - (inverse (exp b) + b / exp b) = integral\<^isup>L lborel (?f b)"
+ finally show "1 - (inverse (exp b) + b / exp b) = integral\<^sup>L lborel (?f b)"
by (auto simp: field_simps exp_minus inverse_eq_divide)
qed
- then have "((\<lambda>i. integral\<^isup>L lborel (?f i)) ---> 1 - (0 + 0)) at_top"
+ then have "((\<lambda>i. integral\<^sup>L lborel (?f i)) ---> 1 - (0 + 0)) at_top"
proof (rule Lim_transform_eventually)
show "((\<lambda>i. 1 - (inverse (exp i) + i / exp i)) ---> 1 - (0 + 0 :: real)) at_top"
using tendsto_power_div_exp_0[of 1]
by (intro tendsto_intros tendsto_inverse_0_at_top exp_at_top) simp
qed
- then show "(\<lambda>i. integral\<^isup>L lborel (?f i)) ----> 1"
+ then show "(\<lambda>i. integral\<^sup>L lborel (?f i)) ----> 1"
by (intro filterlim_compose[OF _ filterlim_real_sequentially]) simp
show "AE x in lborel. mono (\<lambda>n::nat. x * exp (- x) * indicator {0..real n} x)"
by (auto simp: mono_def mult_le_0_iff zero_le_mult_iff split: split_indicator)
@@ -237,10 +237,10 @@
by (simp add: distr[of B] measurable_sets)
also have "\<dots> = (1 / emeasure M' A) * emeasure M' (A \<inter> B)"
by simp
- also have "\<dots> = (\<integral>\<^isup>+ x. (1 / emeasure M' A) * indicator (A \<inter> B) x \<partial>M')"
+ also have "\<dots> = (\<integral>\<^sup>+ x. (1 / emeasure M' A) * indicator (A \<inter> B) x \<partial>M')"
using A B
by (intro positive_integral_cmult_indicator[symmetric]) (auto intro!: zero_le_divide_ereal)
- also have "\<dots> = (\<integral>\<^isup>+ x. ?f x * indicator B x \<partial>M')"
+ also have "\<dots> = (\<integral>\<^sup>+ x. ?f x * indicator B x \<partial>M')"
by (rule positive_integral_cong) (auto split: split_indicator)
finally show "emeasure (distr M M' X) B = emeasure (density M' ?f) B"
using A B X by (auto simp add: emeasure_distr emeasure_density)
@@ -267,15 +267,15 @@
then show "emeasure M {x\<in>space M. X x \<le> a} = ereal (measure lborel (A \<inter> {..a}) / r)"
using distr by simp
- have "(\<integral>\<^isup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
- (\<integral>\<^isup>+ x. ereal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
+ have "(\<integral>\<^sup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
+ (\<integral>\<^sup>+ x. ereal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
by (auto intro!: positive_integral_cong split: split_indicator)
also have "\<dots> = ereal (1 / measure lborel A) * emeasure lborel (A \<inter> {..a})"
using `A \<in> sets borel`
by (intro positive_integral_cmult_indicator) (auto simp: measure_nonneg)
also have "\<dots> = ereal (measure lborel (A \<inter> {..a}) / r)"
unfolding emeasure_eq_ereal_measure[OF fin] using A by (simp add: measure_def)
- finally show "(\<integral>\<^isup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
+ finally show "(\<integral>\<^sup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
ereal (measure lborel (A \<inter> {..a}) / r)" .
qed (auto intro!: divide_nonneg_nonneg measure_nonneg)
@@ -325,7 +325,7 @@
have "emeasure M {x \<in> space M. X x \<le> t} = emeasure (distr M lborel X) {.. t}"
using distributed_measurable[OF D]
by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure])
- also have "\<dots> = (\<integral>\<^isup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
+ also have "\<dots> = (\<integral>\<^sup>+x. ereal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
using distributed_borel_measurable[OF D] `a \<le> t` `t \<le> b`
unfolding distributed_distr_eq_density[OF D]
by (subst emeasure_density)
@@ -378,11 +378,11 @@
show "isCont (\<lambda>x. x / Sigma_Algebra.measure lborel {a..b}) x"
using uniform_distributed_params[OF D]
by (auto intro!: isCont_divide)
- have *: "b\<twosuperior> / (2 * measure lborel {a..b}) - a\<twosuperior> / (2 * measure lborel {a..b}) =
+ have *: "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) =
(b*b - a * a) / (2 * (b - a))"
using `a < b`
by (auto simp: measure_def power2_eq_square diff_divide_distrib[symmetric])
- show "b\<twosuperior> / (2 * measure lborel {a..b}) - a\<twosuperior> / (2 * measure lborel {a..b}) = (a + b) / 2"
+ show "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (a + b) / 2"
using `a < b`
unfolding * square_diff_square_factored by (auto simp: field_simps)
qed (insert `a < b`, simp)
--- a/src/HOL/Probability/Fin_Map.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Fin_Map.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,10 +10,10 @@
text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of
projective limit. @{const extensional} functions are used for the representation in order to
- stay close to the developments of (finite) products @{const Pi\<^isub>E} and their sigma-algebra
- @{const Pi\<^isub>M}. *}
+ stay close to the developments of (finite) products @{const Pi\<^sub>E} and their sigma-algebra
+ @{const Pi\<^sub>M}. *}
-typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^isub>F /_)" [22, 21] 21) =
+typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^sub>F /_)" [22, 21] 21) =
"{(I::'i set, f::'i \<Rightarrow> 'a). finite I \<and> f \<in> extensional I}" by auto
subsection {* Domain and Application *}
@@ -23,11 +23,11 @@
lemma finite_domain[simp, intro]: "finite (domain P)"
by (cases P) (auto simp: domain_def Abs_finmap_inverse)
-definition proj ("'((_)')\<^isub>F" [0] 1000) where "proj P i = snd (Rep_finmap P) i"
+definition proj ("'((_)')\<^sub>F" [0] 1000) where "proj P i = snd (Rep_finmap P) i"
declare [[coercion proj]]
-lemma extensional_proj[simp, intro]: "(P)\<^isub>F \<in> extensional (domain P)"
+lemma extensional_proj[simp, intro]: "(P)\<^sub>F \<in> extensional (domain P)"
by (cases P) (auto simp: domain_def Abs_finmap_inverse proj_def[abs_def])
lemma proj_undefined[simp, intro]: "i \<notin> domain P \<Longrightarrow> P i = undefined"
@@ -42,9 +42,9 @@
instance finmap :: (countable, countable) countable
proof
- obtain mapper where mapper: "\<And>fm :: 'a \<Rightarrow>\<^isub>F 'b. set (mapper fm) = domain fm"
+ obtain mapper where mapper: "\<And>fm :: 'a \<Rightarrow>\<^sub>F 'b. set (mapper fm) = domain fm"
by (metis finite_list[OF finite_domain])
- have "inj (\<lambda>fm. map (\<lambda>i. (i, (fm)\<^isub>F i)) (mapper fm))" (is "inj ?F")
+ have "inj (\<lambda>fm. map (\<lambda>i. (i, (fm)\<^sub>F i)) (mapper fm))" (is "inj ?F")
proof (rule inj_onI)
fix f1 f2 assume "?F f1 = ?F f2"
then have "map fst (?F f1) = map fst (?F f2)" by simp
@@ -54,7 +54,7 @@
unfolding `mapper f1 = mapper f2` map_eq_conv mapper
by (simp add: finmap_eq_iff)
qed
- then show "\<exists>to_nat :: 'a \<Rightarrow>\<^isub>F 'b \<Rightarrow> nat. inj to_nat"
+ then show "\<exists>to_nat :: 'a \<Rightarrow>\<^sub>F 'b \<Rightarrow> nat. inj to_nat"
by (intro exI[of _ "to_nat \<circ> ?F"] inj_comp) auto
qed
@@ -64,7 +64,7 @@
lemma proj_finmap_of[simp]:
assumes "finite inds"
- shows "(finmap_of inds f)\<^isub>F = restrict f inds"
+ shows "(finmap_of inds f)\<^sub>F = restrict f inds"
using assms
by (auto simp: Abs_finmap_inverse finmap_of_def proj_def)
@@ -86,7 +86,7 @@
proof (rule inj_onI)
fix x y::"'a \<Rightarrow> 'b"
assume "finmap_of K x = finmap_of K y"
- hence "(finmap_of K x)\<^isub>F = (finmap_of K y)\<^isub>F" by simp
+ hence "(finmap_of K x)\<^sub>F = (finmap_of K y)\<^sub>F" by simp
moreover
assume "x \<in> S" "y \<in> S" hence "x \<in> extensional K" "y \<in> extensional K" using assms by auto
ultimately
@@ -97,8 +97,8 @@
text {* This is @{term Pi} for Finite Maps, most of this is copied *}
-definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^isub>F 'a) set" where
- "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^isub>F i \<in> A i) } "
+definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^sub>F 'a) set" where
+ "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^sub>F i \<in> A i) } "
syntax
"_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PI' _:_./ _)" 10)
@@ -145,7 +145,7 @@
lemma Pi'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> Pi' A B \<subseteq> Pi' A C"
by (auto simp: Pi'_def)
-lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^isub>E A B) = proj ` Pi' A B"
+lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^sub>E A B) = proj ` Pi' A B"
apply (auto simp: Pi'_def Pi_def extensional_def)
apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI)
apply auto
@@ -156,7 +156,7 @@
instantiation finmap :: (type, topological_space) topological_space
begin
-definition open_finmap :: "('a \<Rightarrow>\<^isub>F 'b) set \<Rightarrow> bool" where
+definition open_finmap :: "('a \<Rightarrow>\<^sub>F 'b) set \<Rightarrow> bool" where
"open_finmap = generate_topology {Pi' a b|a b. \<forall>i\<in>a. open (b i)}"
lemma open_Pi'I: "(\<And>i. i \<in> I \<Longrightarrow> open (A i)) \<Longrightarrow> open (Pi' I A)"
@@ -192,7 +192,7 @@
using open_restricted_space[of "\<lambda>x. \<not> P x"]
unfolding closed_def by (rule back_subst) auto
-lemma tendsto_proj: "((\<lambda>x. x) ---> a) F \<Longrightarrow> ((\<lambda>x. (x)\<^isub>F i) ---> (a)\<^isub>F i) F"
+lemma tendsto_proj: "((\<lambda>x. x) ---> a) F \<Longrightarrow> ((\<lambda>x. (x)\<^sub>F i) ---> (a)\<^sub>F i) F"
unfolding tendsto_def
proof safe
fix S::"'b set"
@@ -200,17 +200,17 @@
assume "open S" hence "open ?S" by (auto intro!: open_Pi'I)
moreover assume "\<forall>S. open S \<longrightarrow> a \<in> S \<longrightarrow> eventually (\<lambda>x. x \<in> S) F" "a i \<in> S"
ultimately have "eventually (\<lambda>x. x \<in> ?S) F" by auto
- thus "eventually (\<lambda>x. (x)\<^isub>F i \<in> S) F"
+ thus "eventually (\<lambda>x. (x)\<^sub>F i \<in> S) F"
by eventually_elim (insert `a i \<in> S`, force simp: Pi'_iff split: split_if_asm)
qed
lemma continuous_proj:
- shows "continuous_on s (\<lambda>x. (x)\<^isub>F i)"
+ shows "continuous_on s (\<lambda>x. (x)\<^sub>F i)"
unfolding continuous_on_def by (safe intro!: tendsto_proj tendsto_ident_at)
instance finmap :: (type, first_countable_topology) first_countable_topology
proof
- fix x::"'a\<Rightarrow>\<^isub>F'b"
+ fix x::"'a\<Rightarrow>\<^sub>F'b"
have "\<forall>i. \<exists>A. countable A \<and> (\<forall>a\<in>A. x i \<in> a) \<and> (\<forall>a\<in>A. open a) \<and>
(\<forall>S. open S \<and> x i \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)) \<and> (\<forall>a b. a \<in> A \<longrightarrow> b \<in> A \<longrightarrow> a \<inter> b \<in> A)" (is "\<forall>i. ?th i")
proof
@@ -220,19 +220,19 @@
then guess A unfolding choice_iff .. note A = this
hence open_sub: "\<And>i S. i\<in>domain x \<Longrightarrow> open (S i) \<Longrightarrow> x i\<in>(S i) \<Longrightarrow> (\<exists>a\<in>A i. a\<subseteq>(S i))" by auto
have A_notempty: "\<And>i. i \<in> domain x \<Longrightarrow> A i \<noteq> {}" using open_sub[of _ "\<lambda>_. UNIV"] by auto
- let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^isub>E (domain x) A)"
- show "\<exists>A::nat \<Rightarrow> ('a\<Rightarrow>\<^isub>F'b) set. (\<forall>i. x \<in> (A i) \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
+ let ?A = "(\<lambda>f. Pi' (domain x) f) ` (Pi\<^sub>E (domain x) A)"
+ show "\<exists>A::nat \<Rightarrow> ('a\<Rightarrow>\<^sub>F'b) set. (\<forall>i. x \<in> (A i) \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))"
proof (rule first_countableI[where A="?A"], safe)
show "countable ?A" using A by (simp add: countable_PiE)
next
- fix S::"('a \<Rightarrow>\<^isub>F 'b) set" assume "open S" "x \<in> S"
+ fix S::"('a \<Rightarrow>\<^sub>F 'b) set" assume "open S" "x \<in> S"
thus "\<exists>a\<in>?A. a \<subseteq> S" unfolding open_finmap_def
proof (induct rule: generate_topology.induct)
case UNIV thus ?case by (auto simp add: ex_in_conv PiE_eq_empty_iff A_notempty)
next
case (Int a b)
then obtain f g where
- "f \<in> Pi\<^isub>E (domain x) A" "Pi' (domain x) f \<subseteq> a" "g \<in> Pi\<^isub>E (domain x) A" "Pi' (domain x) g \<subseteq> b"
+ "f \<in> Pi\<^sub>E (domain x) A" "Pi' (domain x) f \<subseteq> a" "g \<in> Pi\<^sub>E (domain x) A" "Pi' (domain x) g \<subseteq> b"
by auto
thus ?case using A
by (auto simp: Pi'_iff PiE_iff extensional_def Int_stable_def
@@ -245,10 +245,10 @@
next
case (Basis s)
then obtain a b where xs: "x\<in> Pi' a b" "s = Pi' a b" "\<And>i. i\<in>a \<Longrightarrow> open (b i)" by auto
- have "\<forall>i. \<exists>a. (i \<in> domain x \<and> open (b i) \<and> (x)\<^isub>F i \<in> b i) \<longrightarrow> (a\<in>A i \<and> a \<subseteq> b i)"
+ have "\<forall>i. \<exists>a. (i \<in> domain x \<and> open (b i) \<and> (x)\<^sub>F i \<in> b i) \<longrightarrow> (a\<in>A i \<and> a \<subseteq> b i)"
using open_sub[of _ b] by auto
then obtain b'
- where "\<And>i. i \<in> domain x \<Longrightarrow> open (b i) \<Longrightarrow> (x)\<^isub>F i \<in> b i \<Longrightarrow> (b' i \<in>A i \<and> b' i \<subseteq> b i)"
+ where "\<And>i. i \<in> domain x \<Longrightarrow> open (b i) \<Longrightarrow> (x)\<^sub>F i \<in> b i \<Longrightarrow> (b' i \<in>A i \<and> b' i \<subseteq> b i)"
unfolding choice_iff by auto
with xs have "\<And>i. i \<in> a \<Longrightarrow> (b' i \<in>A i \<and> b' i \<subseteq> b i)" "Pi' a b' \<subseteq> Pi' a b"
by (auto simp: Pi'_iff intro!: Pi'_mono)
@@ -265,7 +265,7 @@
begin
definition dist_finmap where
- "dist P Q = Max (range (\<lambda>i. dist ((P)\<^isub>F i) ((Q)\<^isub>F i))) + (if domain P = domain Q then 0 else 1)"
+ "dist P Q = Max (range (\<lambda>i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i))) + (if domain P = domain Q then 0 else 1)"
lemma add_eq_zero_iff[simp]:
fixes a b::real
@@ -273,16 +273,16 @@
shows "a + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
using assms by auto
-lemma finite_proj_image': "x \<notin> domain P \<Longrightarrow> finite ((P)\<^isub>F ` S)"
+lemma finite_proj_image': "x \<notin> domain P \<Longrightarrow> finite ((P)\<^sub>F ` S)"
by (rule finite_subset[of _ "proj P ` (domain P \<inter> S \<union> {x})"]) auto
-lemma finite_proj_image: "finite ((P)\<^isub>F ` S)"
+lemma finite_proj_image: "finite ((P)\<^sub>F ` S)"
by (cases "\<exists>x. x \<notin> domain P") (auto intro: finite_proj_image' finite_subset[where B="domain P"])
-lemma finite_proj_diag: "finite ((\<lambda>i. d ((P)\<^isub>F i) ((Q)\<^isub>F i)) ` S)"
+lemma finite_proj_diag: "finite ((\<lambda>i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S)"
proof -
- have "(\<lambda>i. d ((P)\<^isub>F i) ((Q)\<^isub>F i)) ` S = (\<lambda>(i, j). d i j) ` ((\<lambda>i. ((P)\<^isub>F i, (Q)\<^isub>F i)) ` S)" by auto
- moreover have "((\<lambda>i. ((P)\<^isub>F i, (Q)\<^isub>F i)) ` S) \<subseteq> (\<lambda>i. (P)\<^isub>F i) ` S \<times> (\<lambda>i. (Q)\<^isub>F i) ` S" by auto
+ have "(\<lambda>i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S = (\<lambda>(i, j). d i j) ` ((\<lambda>i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S)" by auto
+ moreover have "((\<lambda>i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S) \<subseteq> (\<lambda>i. (P)\<^sub>F i) ` S \<times> (\<lambda>i. (Q)\<^sub>F i) ` S" by auto
moreover have "finite \<dots>" using finite_proj_image[of P S] finite_proj_image[of Q S]
by (intro finite_cartesian_product) simp_all
ultimately show ?thesis by (simp add: finite_subset)
@@ -293,7 +293,7 @@
by (simp add: dist_finmap_def finite_proj_diag split: split_if_asm)
lemma dist_proj:
- shows "dist ((x)\<^isub>F i) ((y)\<^isub>F i) \<le> dist x y"
+ shows "dist ((x)\<^sub>F i) ((y)\<^sub>F i) \<le> dist x y"
proof -
have "dist (x i) (y i) \<le> Max (range (\<lambda>i. dist (x i) (y i)))"
by (simp add: Max_ge_iff finite_proj_diag)
@@ -312,7 +312,7 @@
also have "\<dots> < e"
proof (subst Max_less_iff, safe)
fix i
- show "dist ((P)\<^isub>F i) ((Q)\<^isub>F i) < e" using assms
+ show "dist ((P)\<^sub>F i) ((Q)\<^sub>F i) < e" using assms
by (cases "i \<in> domain P") simp_all
qed (simp add: finite_proj_diag)
finally show ?thesis .
@@ -320,7 +320,7 @@
instance
proof
- fix S::"('a \<Rightarrow>\<^isub>F 'b) set"
+ fix S::"('a \<Rightarrow>\<^sub>F 'b) set"
show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" (is "_ = ?od")
proof
assume "open S"
@@ -365,7 +365,7 @@
show "domain y = a" using d s `a \<noteq> {}` by (auto simp: dist_le_1_imp_domain_eq a_dom)
fix i assume "i \<in> a"
moreover
- hence "dist ((y)\<^isub>F i) ((x)\<^isub>F i) < es i" using d
+ hence "dist ((y)\<^sub>F i) ((x)\<^sub>F i) < es i" using d
by (auto simp: dist_finmap_def `a \<noteq> {}` intro!: le_less_trans[OF dist_proj])
ultimately
show "y i \<in> b i" by (rule in_b)
@@ -403,15 +403,15 @@
finally show "open S" .
qed
next
- fix P Q::"'a \<Rightarrow>\<^isub>F 'b"
+ fix P Q::"'a \<Rightarrow>\<^sub>F 'b"
have Max_eq_iff: "\<And>A m. finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> (Max A = m) = (m \<in> A \<and> (\<forall>a\<in>A. a \<le> m))"
by (auto intro: Max_in Max_eqI)
show "dist P Q = 0 \<longleftrightarrow> P = Q"
by (auto simp: finmap_eq_iff dist_finmap_def Max_ge_iff finite_proj_diag Max_eq_iff
intro!: Max_eqI image_eqI[where x=undefined])
next
- fix P Q R::"'a \<Rightarrow>\<^isub>F 'b"
- let ?dists = "\<lambda>P Q i. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)"
+ fix P Q R::"'a \<Rightarrow>\<^sub>F 'b"
+ let ?dists = "\<lambda>P Q i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i)"
let ?dpq = "?dists P Q" and ?dpr = "?dists P R" and ?dqr = "?dists Q R"
let ?dom = "\<lambda>P Q. (if domain P = domain Q then 0 else 1::real)"
have "dist P Q = Max (range ?dpq) + ?dom P Q"
@@ -430,21 +430,21 @@
subsection {* Complete Space of Finite Maps *}
lemma tendsto_finmap:
- fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^isub>F ('a::metric_space))"
+ fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^sub>F ('a::metric_space))"
assumes ind_f: "\<And>n. domain (f n) = domain g"
assumes proj_g: "\<And>i. i \<in> domain g \<Longrightarrow> (\<lambda>n. (f n) i) ----> g i"
shows "f ----> g"
unfolding tendsto_iff
proof safe
fix e::real assume "0 < e"
- let ?dists = "\<lambda>x i. dist ((f x)\<^isub>F i) ((g)\<^isub>F i)"
+ let ?dists = "\<lambda>x i. dist ((f x)\<^sub>F i) ((g)\<^sub>F i)"
have "eventually (\<lambda>x. \<forall>i\<in>domain g. ?dists x i < e) sequentially"
using finite_domain[of g] proj_g
proof induct
case (insert i G)
with `0 < e` have "eventually (\<lambda>x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff)
moreover
- from insert have "eventually (\<lambda>x. \<forall>i\<in>G. dist ((f x)\<^isub>F i) ((g)\<^isub>F i) < e) sequentially" by simp
+ from insert have "eventually (\<lambda>x. \<forall>i\<in>G. dist ((f x)\<^sub>F i) ((g)\<^sub>F i) < e) sequentially" by simp
ultimately show ?case by eventually_elim auto
qed simp
thus "eventually (\<lambda>x. dist (f x) g < e) sequentially"
@@ -453,7 +453,7 @@
instance finmap :: (type, complete_space) complete_space
proof
- fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^isub>F 'b"
+ fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^sub>F 'b"
assume "Cauchy P"
then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
by (force simp: cauchy)
@@ -507,7 +507,7 @@
assume "i \<in> domain (P n)"
hence "ni i \<le> Max (ni ` d)" using dom by simp
also have "\<dots> \<le> N" by (simp add: N_def)
- finally show "dist ((P n)\<^isub>F i) ((Q)\<^isub>F i) < e" using ni `i \<in> domain (P n)` `N \<le> n` dom
+ finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni `i \<in> domain (P n)` `N \<le> n` dom
by (auto simp: p_def q N_def less_imp_le)
qed
qed
@@ -526,7 +526,7 @@
lemma countable_basis_proj: "countable basis_proj" and basis_proj: "topological_basis basis_proj"
unfolding basis_proj_def by (intro is_basis countable_basis)+
-definition basis_finmap::"('a \<Rightarrow>\<^isub>F 'b) set set"
+definition basis_finmap::"('a \<Rightarrow>\<^sub>F 'b) set set"
where "basis_finmap = {Pi' I S|I S. finite I \<and> (\<forall>i \<in> I. S i \<in> basis_proj)}"
lemma in_basis_finmapI:
@@ -536,8 +536,8 @@
lemma basis_finmap_eq:
assumes "basis_proj \<noteq> {}"
- shows "basis_finmap = (\<lambda>f. Pi' (domain f) (\<lambda>i. from_nat_into basis_proj ((f)\<^isub>F i))) `
- (UNIV::('a \<Rightarrow>\<^isub>F nat) set)" (is "_ = ?f ` _")
+ shows "basis_finmap = (\<lambda>f. Pi' (domain f) (\<lambda>i. from_nat_into basis_proj ((f)\<^sub>F i))) `
+ (UNIV::('a \<Rightarrow>\<^sub>F nat) set)" (is "_ = ?f ` _")
unfolding basis_finmap_def
proof safe
fix I::"'a set" and S::"'a \<Rightarrow> 'b set"
@@ -546,8 +546,8 @@
by (force simp: Pi'_def countable_basis_proj)
thus "Pi' I S \<in> range ?f" by simp
next
- fix x and f::"'a \<Rightarrow>\<^isub>F nat"
- show "\<exists>I S. (\<Pi>' i\<in>domain f. from_nat_into local.basis_proj ((f)\<^isub>F i)) = Pi' I S \<and>
+ fix x and f::"'a \<Rightarrow>\<^sub>F nat"
+ show "\<exists>I S. (\<Pi>' i\<in>domain f. from_nat_into local.basis_proj ((f)\<^sub>F i)) = Pi' I S \<and>
finite I \<and> (\<forall>i\<in>I. S i \<in> local.basis_proj)"
using assms by (auto intro: from_nat_into)
qed
@@ -566,7 +566,7 @@
by (auto intro!: open_Pi'I topological_basis_open[OF basis_proj]
simp: topological_basis_def basis_finmap_def Let_def)
next
- fix O'::"('a \<Rightarrow>\<^isub>F 'b) set" and x
+ fix O'::"('a \<Rightarrow>\<^sub>F 'b) set" and x
assume O': "open O'" "x \<in> O'"
then obtain a where a:
"x \<in> Pi' (domain x) a" "Pi' (domain x) a \<subseteq> O'" "\<And>i. i\<in>domain x \<Longrightarrow> open (a i)"
@@ -619,16 +619,16 @@
sigma (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j))) {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
abbreviation
- "Pi\<^isub>F I M \<equiv> PiF I M"
+ "Pi\<^sub>F I M \<equiv> PiF I M"
syntax
"_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3PIF _:_./ _)" 10)
syntax (xsymbols)
- "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>F _\<in>_./ _)" 10)
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>F _\<in>_./ _)" 10)
syntax (HTML output)
- "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>F _\<in>_./ _)" 10)
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>F _\<in>_./ _)" 10)
translations
"PIF x:I. M" == "CONST PiF I (%x. M)"
@@ -761,7 +761,7 @@
proof safe
fix y assume "y \<in> sets N"
have "A -` y = (\<Union>{A -` y \<inter> {x. domain x = J}|J. finite J})" by auto
- { fix x::"'a \<Rightarrow>\<^isub>F 'b"
+ { fix x::"'a \<Rightarrow>\<^sub>F 'b"
from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto
hence "\<exists>n. domain x = set (from_nat n)"
by (intro exI[where x="to_nat xs"]) auto }
@@ -860,7 +860,7 @@
lemma measurable_PiM_finmap_of:
assumes "finite J"
- shows "finmap_of J \<in> measurable (Pi\<^isub>M J M) (PiF {J} M)"
+ shows "finmap_of J \<in> measurable (Pi\<^sub>M J M) (PiF {J} M)"
apply (rule measurable_finmap_of)
apply (rule measurable_component_singleton)
apply simp
@@ -871,10 +871,10 @@
lemma proj_measurable_singleton:
assumes "A \<in> sets (M i)"
- shows "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
+ shows "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
proof cases
assume "i \<in> I"
- hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
+ hence "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) =
Pi' I (\<lambda>x. if x = i then A else space (M x))"
using sets.sets_into_space[OF ] `A \<in> sets (M i)` assms
by (auto simp: space_PiF Pi'_def)
@@ -882,21 +882,21 @@
by (intro in_sets_PiFI) auto
next
assume "i \<notin> I"
- hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
+ hence "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space (PiF {I} M) =
(if undefined \<in> A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def)
thus ?thesis by simp
qed
lemma measurable_proj_singleton:
assumes "i \<in> I"
- shows "(\<lambda>x. (x)\<^isub>F i) \<in> measurable (PiF {I} M) (M i)"
+ shows "(\<lambda>x. (x)\<^sub>F i) \<in> measurable (PiF {I} M) (M i)"
by (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
(insert `i \<in> I`, auto simp: space_PiF)
lemma measurable_proj_countable:
fixes I::"'a::countable set set"
assumes "y \<in> space (M i)"
- shows "(\<lambda>x. if i \<in> domain x then (x)\<^isub>F i else y) \<in> measurable (PiF I M) (M i)"
+ shows "(\<lambda>x. if i \<in> domain x then (x)\<^sub>F i else y) \<in> measurable (PiF I M) (M i)"
proof (rule countable_measurable_PiFI)
fix J assume "J \<in> I" "finite J"
show "(\<lambda>x. if i \<in> domain x then x i else y) \<in> measurable (PiF {J} M) (M i)"
@@ -904,7 +904,7 @@
proof safe
fix z assume "z \<in> sets (M i)"
have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
- (\<lambda>x. if i \<in> J then (x)\<^isub>F i else y) -` z \<inter> space (PiF {J} M)"
+ (\<lambda>x. if i \<in> J then (x)\<^sub>F i else y) -` z \<inter> space (PiF {J} M)"
by (auto simp: space_PiF Pi'_def)
also have "\<dots> \<in> sets (PiF {J} M)" using `z \<in> sets (M i)` `finite J`
by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
@@ -925,14 +925,14 @@
assumes "x \<in> space (PiM J M)"
shows "proj \<in> measurable (PiF {J} M) (PiM J M)"
proof (rule measurable_PiM_single)
- show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^isub>E i \<in> J. space (M i))"
+ show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^sub>E i \<in> J. space (M i))"
using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def)
next
fix A i assume A: "i \<in> J" "A \<in> sets (M i)"
- show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} \<in> sets (PiF {J} M)"
+ show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^sub>F i \<in> A} \<in> sets (PiF {J} M)"
proof
- have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} =
- (\<lambda>\<omega>. (\<omega>)\<^isub>F i) -` A \<inter> space (PiF {J} M)" by auto
+ have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^sub>F i \<in> A} =
+ (\<lambda>\<omega>. (\<omega>)\<^sub>F i) -` A \<inter> space (PiF {J} M)" by auto
also have "\<dots> \<in> sets (PiF {J} M)"
using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM)
finally show ?thesis .
@@ -1024,11 +1024,11 @@
defines "P == { Pi' I F | F. \<forall>i\<in>I. F i \<in> E i }"
shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P"
proof
- let ?P = "sigma (space (Pi\<^isub>F {I} M)) P"
+ let ?P = "sigma (space (Pi\<^sub>F {I} M)) P"
from `finite I`[THEN ex_bij_betw_finite_nat] guess T ..
then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i"
by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f simp del: `finite I`)
- have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>F {I} M))"
+ have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>F {I} M))"
using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq)
then have space_P: "space ?P = (\<Pi>' i\<in>I. space (M i))"
by (simp add: space_PiF)
@@ -1038,15 +1038,15 @@
also have "\<dots> \<subseteq> sets (sigma (space (PiF {I} M)) P)"
proof (safe intro!: sets.sigma_sets_subset)
fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
- have "(\<lambda>x. (x)\<^isub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
+ have "(\<lambda>x. (x)\<^sub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
proof (subst measurable_iff_measure_of)
show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
- from space_P `i \<in> I` show "(\<lambda>x. (x)\<^isub>F i) \<in> space ?P \<rightarrow> space (M i)"
+ from space_P `i \<in> I` show "(\<lambda>x. (x)\<^sub>F i) \<in> space ?P \<rightarrow> space (M i)"
by auto
- show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
proof
fix A assume A: "A \<in> E i"
- then have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
+ then have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
by (intro Pi'_cong) (simp_all add: S_union)
@@ -1065,14 +1065,14 @@
by (simp add: P_closed)
(auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"])
qed
- finally show "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ finally show "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
using P_closed by simp
qed
qed
from measurable_sets[OF this, of A] A `i \<in> I` E_closed
- have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
+ have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P \<in> sets ?P"
by (simp add: E_generates)
- also have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
+ also have "(\<lambda>x. (x)\<^sub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
using P_closed by (auto simp: space_PiF)
finally show "\<dots> \<in> sets ?P" .
qed
@@ -1114,11 +1114,11 @@
lemma finmap_UNIV[simp]: "(\<Union>J\<in>Collect finite. PI' j : J. UNIV) = UNIV" by auto
lemma borel_eq_PiF_borel:
- shows "(borel :: ('i::countable \<Rightarrow>\<^isub>F 'a::polish_space) measure) =
+ shows "(borel :: ('i::countable \<Rightarrow>\<^sub>F 'a::polish_space) measure) =
PiF (Collect finite) (\<lambda>_. borel :: 'a measure)"
unfolding borel_def PiF_def
proof (rule measure_eqI, clarsimp, rule sigma_sets_eqI)
- fix a::"('i \<Rightarrow>\<^isub>F 'a) set" assume "a \<in> Collect open" hence "open a" by simp
+ fix a::"('i \<Rightarrow>\<^sub>F 'a) set" assume "a \<in> Collect open" hence "open a" by simp
then obtain B' where B': "B'\<subseteq>basis_finmap" "a = \<Union>B'"
using finmap_topological_basis by (force simp add: topological_basis_def)
have "a \<in> sigma UNIV {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
@@ -1138,9 +1138,9 @@
thus "a \<in> sigma_sets UNIV {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
by simp
next
- fix b::"('i \<Rightarrow>\<^isub>F 'a) set"
+ fix b::"('i \<Rightarrow>\<^sub>F 'a) set"
assume "b \<in> {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
- hence b': "b \<in> sets (Pi\<^isub>F (Collect finite) (\<lambda>_. borel))" by (auto simp: sets_PiF borel_def)
+ hence b': "b \<in> sets (Pi\<^sub>F (Collect finite) (\<lambda>_. borel))" by (auto simp: sets_PiF borel_def)
let ?b = "\<lambda>J. b \<inter> {x. domain x = J}"
have "b = \<Union>((\<lambda>J. ?b J) ` Collect finite)" by auto
also have "\<dots> \<in> sets borel"
@@ -1202,16 +1202,16 @@
lemma fm_product:
assumes "\<And>i. space (M i) = UNIV"
- shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^isub>M J M) = (\<Pi>\<^isub>E j \<in> J. S (f j))"
+ shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^sub>M J M) = (\<Pi>\<^sub>E j \<in> J. S (f j))"
using assms
by (auto simp: inv fm_def compose_def space_PiM Pi'_def)
lemma fm_measurable:
assumes "f ` J \<in> N"
- shows "fm \<in> measurable (Pi\<^isub>M J (\<lambda>_. M)) (Pi\<^isub>F N (\<lambda>_. M))"
+ shows "fm \<in> measurable (Pi\<^sub>M J (\<lambda>_. M)) (Pi\<^sub>F N (\<lambda>_. M))"
unfolding fm_def
proof (rule measurable_comp, rule measurable_compose_inv)
- show "finmap_of (f ` J) \<in> measurable (Pi\<^isub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
+ show "finmap_of (f ` J) \<in> measurable (Pi\<^sub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
using assms by (intro measurable_finmap_of measurable_component_singleton) auto
qed (simp_all add: inv)
@@ -1229,7 +1229,7 @@
lemma inj_on_fm:
assumes "\<And>i. space (M i) = UNIV"
- shows "inj_on fm (space (Pi\<^isub>M J M))"
+ shows "inj_on fm (space (Pi\<^sub>M J M))"
using assms
apply (auto simp: fm_def space_PiM PiE_def)
apply (rule comp_inj_on)
@@ -1244,7 +1244,7 @@
definition "mf = (\<lambda>g. compose J g f) o proj"
lemma mf_fm:
- assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))"
+ assumes "x \<in> space (Pi\<^sub>M J (\<lambda>_. M))"
shows "mf (fm x) = x"
proof -
have "mf (fm x) \<in> extensional J"
@@ -1266,13 +1266,13 @@
shows "mf \<in> measurable (PiF {f ` J} (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
unfolding mf_def
proof (rule measurable_comp, rule measurable_proj_PiM)
- show "(\<lambda>g. compose J g f) \<in> measurable (Pi\<^isub>M (f ` J) (\<lambda>x. M)) (Pi\<^isub>M J (\<lambda>_. M))"
+ show "(\<lambda>g. compose J g f) \<in> measurable (Pi\<^sub>M (f ` J) (\<lambda>x. M)) (Pi\<^sub>M J (\<lambda>_. M))"
by (rule measurable_finmap_compose)
qed (auto simp add: space_PiM extensional_def assms)
lemma fm_image_measurable:
assumes "space M = UNIV"
- assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M))"
+ assumes "X \<in> sets (Pi\<^sub>M J (\<lambda>_. M))"
shows "fm ` X \<in> sets (PiF {f ` J} (\<lambda>_. M))"
proof -
have "fm ` X = (mf) -` X \<inter> space (PiF {f ` J} (\<lambda>_. M))"
@@ -1296,7 +1296,7 @@
lemma fm_image_measurable_finite:
assumes "space M = UNIV"
- assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M::'c measure))"
+ assumes "X \<in> sets (Pi\<^sub>M J (\<lambda>_. M::'c measure))"
shows "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. M::'c measure))"
using fm_image_measurable[OF assms]
by (rule subspace_set_in_sets) (auto simp: finite_subset)
@@ -1312,8 +1312,8 @@
unfolding mapmeasure_def by simp
lemma mapmeasure_PiF:
- assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
- assumes s2: "sets M = sets (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes s1: "space M = space (Pi\<^sub>M J (\<lambda>_. N))"
+ assumes s2: "sets M = sets (Pi\<^sub>M J (\<lambda>_. N))"
assumes "space N = UNIV"
assumes "X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
shows "emeasure (mapmeasure M (\<lambda>_. N)) X = emeasure M ((fm -` X \<inter> extensional J))"
@@ -1323,15 +1323,15 @@
lemma mapmeasure_PiM:
fixes N::"'c measure"
- assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
- assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
+ assumes s1: "space M = space (Pi\<^sub>M J (\<lambda>_. N))"
+ assumes s2: "sets M = (Pi\<^sub>M J (\<lambda>_. N))"
assumes N: "space N = UNIV"
assumes X: "X \<in> sets M"
shows "emeasure M X = emeasure (mapmeasure M (\<lambda>_. N)) (fm ` X)"
unfolding mapmeasure_def
proof (subst emeasure_distr, subst measurable_eqI[OF s1 refl s2 refl], rule fm_measurable)
- have "X \<subseteq> space (Pi\<^isub>M J (\<lambda>_. N))" using assms by (simp add: sets.sets_into_space)
- from assms inj_on_fm[of "\<lambda>_. N"] set_mp[OF this] have "fm -` fm ` X \<inter> space (Pi\<^isub>M J (\<lambda>_. N)) = X"
+ have "X \<subseteq> space (Pi\<^sub>M J (\<lambda>_. N))" using assms by (simp add: sets.sets_into_space)
+ from assms inj_on_fm[of "\<lambda>_. N"] set_mp[OF this] have "fm -` fm ` X \<inter> space (Pi\<^sub>M J (\<lambda>_. N)) = X"
by (auto simp: vimage_image_eq inj_on_def)
thus "emeasure M X = emeasure M (fm -` fm ` X \<inter> space M)" using s1
by simp
--- a/src/HOL/Probability/Finite_Product_Measure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -77,8 +77,8 @@
lemma injective_vimage_restrict:
assumes J: "J \<subseteq> I"
- and sets: "A \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" and ne: "(\<Pi>\<^isub>E i\<in>I. S i) \<noteq> {}"
- and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
+ and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}"
+ and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
shows "A = B"
proof (intro set_eqI)
fix x
@@ -86,8 +86,8 @@
have "J \<inter> (I - J) = {}" by auto
show "x \<in> A \<longleftrightarrow> x \<in> B"
proof cases
- assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"
- have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
+ assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
+ have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S]
by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
then show "x \<in> A \<longleftrightarrow> x \<in> B"
@@ -98,11 +98,11 @@
lemma restrict_vimage:
"I \<inter> J = {} \<Longrightarrow>
- (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
+ (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
by (auto simp: restrict_Pi_cancel PiE_def)
lemma merge_vimage:
- "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
+ "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
by (auto simp: restrict_Pi_cancel PiE_def)
section "Finite product spaces"
@@ -126,11 +126,11 @@
by (auto simp: prod_emb_def)
lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
- prod_emb I M J (\<Pi>\<^isub>E i\<in>J. E i) = (\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i))"
+ prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
by (force simp: prod_emb_def PiE_iff split_if_mem2)
lemma prod_emb_PiE_same_index[simp]:
- "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^isub>E I E) = Pi\<^isub>E I E"
+ "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
by (auto simp: prod_emb_def PiE_iff)
lemma prod_emb_trans[simp]:
@@ -139,12 +139,12 @@
lemma prod_emb_Pi:
assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
- shows "prod_emb K M J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
+ shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
using assms sets.space_closed
by (auto simp: prod_emb_def PiE_iff split: split_if_asm) blast+
lemma prod_emb_id:
- "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
+ "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
by (auto simp: prod_emb_def subset_eq extensional_restrict)
lemma prod_emb_mono:
@@ -152,51 +152,51 @@
by (auto simp: prod_emb_def)
definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
- "PiM I M = extend_measure (\<Pi>\<^isub>E i\<in>I. space (M i))
+ "PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
- (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j))
+ (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
(\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
- "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j)) `
+ "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
abbreviation
- "Pi\<^isub>M I M \<equiv> PiM I M"
+ "Pi\<^sub>M I M \<equiv> PiM I M"
syntax
"_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3PIM _:_./ _)" 10)
syntax (xsymbols)
- "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10)
+ "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10)
syntax (HTML output)
- "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10)
+ "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10)
translations
"PIM x:I. M" == "CONST PiM I (%x. M)"
lemma prod_algebra_sets_into_space:
- "prod_algebra I M \<subseteq> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
by (auto simp: prod_emb_def prod_algebra_def)
lemma prod_algebra_eq_finite:
assumes I: "finite I"
- shows "prod_algebra I M = {(\<Pi>\<^isub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
+ shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
proof (intro iffI set_eqI)
fix A assume "A \<in> ?L"
then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
and A: "A = prod_emb I M J (PIE j:J. E j)"
by (auto simp: prod_algebra_def)
- let ?A = "\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i)"
+ let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)"
have A: "A = ?A"
unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto
show "A \<in> ?R" unfolding A using J sets.top
by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp
next
fix A assume "A \<in> ?R"
- then obtain X where A: "A = (\<Pi>\<^isub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
- then have A: "A = prod_emb I M I (\<Pi>\<^isub>E i\<in>I. X i)"
+ then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
+ then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)"
by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff)
from X I show "A \<in> ?L" unfolding A
by (auto simp: prod_algebra_def)
@@ -208,13 +208,13 @@
by (auto simp: prod_algebra_def)
lemma prod_algebraI_finite:
- "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^isub>E I E) \<in> prod_algebra I M"
+ "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
-lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
+lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
proof (safe intro!: Int_stableI)
fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
- then show "\<exists>G. Pi\<^isub>E J E \<inter> Pi\<^isub>E J F = Pi\<^isub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
+ then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
qed
@@ -226,14 +226,14 @@
lemma prod_algebraE_all:
assumes A: "A \<in> prod_algebra I M"
- obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
+ obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
proof -
- from A obtain E J where A: "A = prod_emb I M J (Pi\<^isub>E J E)"
+ from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
by (auto simp: prod_algebra_def)
from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
using sets.sets_into_space by auto
- then have "A = (\<Pi>\<^isub>E i\<in>I. if i\<in>J then E i else space (M i))"
+ then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))"
using A J by (auto simp: prod_emb_PiE)
moreover then have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
using sets.top E by auto
@@ -246,7 +246,7 @@
from prod_algebraE[OF this] guess J E . note A = this
fix B assume "B \<in> prod_algebra I M"
from prod_algebraE[OF this] guess K F . note B = this
- have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^isub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
+ have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
(if i \<in> K then F i else space (M i)))"
unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4)
B(5)[THEN sets.sets_into_space]
@@ -268,13 +268,13 @@
proof
fix A assume "A \<in> prod_algebra I E"
then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
- and A: "A = prod_emb I E J (\<Pi>\<^isub>E i\<in>J. G i)"
+ and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
by (auto simp: prod_algebra_def)
moreover
- from space have "(\<Pi>\<^isub>E i\<in>I. space (E i)) = (\<Pi>\<^isub>E i\<in>I. space (F i))"
+ from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))"
by (rule PiE_cong)
- with A have "A = prod_emb I F J (\<Pi>\<^isub>E i\<in>J. G i)"
+ with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)"
by (simp add: prod_emb_def)
moreover
from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)"
@@ -297,28 +297,28 @@
qed
lemma space_in_prod_algebra:
- "(\<Pi>\<^isub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
+ "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
proof cases
assume "I = {}" then show ?thesis
by (auto simp add: prod_algebra_def image_iff prod_emb_def)
next
assume "I \<noteq> {}"
then obtain i where "i \<in> I" by auto
- then have "(\<Pi>\<^isub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i))"
+ then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
by (auto simp: prod_emb_def)
also have "\<dots> \<in> prod_algebra I M"
using `i \<in> I` by (intro prod_algebraI) auto
finally show ?thesis .
qed
-lemma space_PiM: "space (\<Pi>\<^isub>M i\<in>I. M i) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
-lemma sets_PiM: "sets (\<Pi>\<^isub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"
+lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
lemma sets_PiM_single: "sets (PiM I M) =
- sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
+ sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
(is "_ = sigma_sets ?\<Omega> ?R")
unfolding sets_PiM
proof (rule sigma_sets_eqI)
@@ -332,7 +332,7 @@
with `I = {}` show ?thesis by (auto intro!: sigma_sets_top)
next
assume "I \<noteq> {}"
- with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^isub>E i\<in>I. space (M i)). f j \<in> X j})"
+ with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
by (auto simp: prod_emb_def)
also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
using X `I \<noteq> {}` by (intro R.finite_INT sigma_sets.Basic) auto
@@ -340,9 +340,9 @@
qed
next
fix A assume "A \<in> ?R"
- then obtain i B where A: "A = {f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
+ then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
by auto
- then have "A = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. B)"
+ then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)"
by (auto simp: prod_emb_def)
also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
using A by (intro sigma_sets.Basic prod_algebraI) auto
@@ -364,9 +364,9 @@
qed
lemma measurable_PiM:
- assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
- f -` prod_emb I M J (Pi\<^isub>E J X) \<inter> space N \<in> sets N"
+ f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
shows "f \<in> measurable N (PiM I M)"
using sets_PiM prod_algebra_sets_into_space space
proof (rule measurable_sigma_sets)
@@ -376,7 +376,7 @@
qed
lemma measurable_PiM_Collect:
- assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
{\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
shows "f \<in> measurable N (PiM I M)"
@@ -391,13 +391,13 @@
qed
lemma measurable_PiM_single:
- assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N"
shows "f \<in> measurable N (PiM I M)"
using sets_PiM_single
proof (rule measurable_sigma_sets)
- fix A assume "A \<in> {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
- then obtain B i where "A = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
+ fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+ then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
by auto
with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto
also have "\<dots> \<in> sets N" using B by (rule sets)
@@ -406,8 +406,8 @@
lemma measurable_PiM_single':
assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
- and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
- shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^isub>M I M)"
+ and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+ shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)"
proof (rule measurable_PiM_single)
fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N"
@@ -422,18 +422,18 @@
using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] `finite I` sets by auto
lemma measurable_component_singleton:
- assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
+ assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)"
proof (unfold measurable_def, intro CollectI conjI ballI)
fix A assume "A \<in> sets (M i)"
- then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = prod_emb I M {i} (\<Pi>\<^isub>E j\<in>{i}. A)"
+ then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
using sets.sets_into_space `i \<in> I`
by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm)
- then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
+ then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
using `A \<in> sets (M i)` `i \<in> I` by (auto intro!: sets_PiM_I)
qed (insert `i \<in> I`, auto simp: space_PiM)
lemma measurable_component_singleton'[measurable_app]:
- assumes f: "f \<in> measurable N (Pi\<^isub>M I M)"
+ assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
assumes i: "i \<in> I"
shows "(\<lambda>x. (f x) i) \<in> measurable N (M i)"
using measurable_compose[OF f measurable_component_singleton, OF i] .
@@ -449,18 +449,18 @@
by (cases i) simp_all
lemma measurable_nat_case'[measurable (raw)]:
- assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
- shows "(\<lambda>x. nat_case (f x) (g x)) \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
+ assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
+ shows "(\<lambda>x. nat_case (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
using fg[THEN measurable_space]
by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
lemma measurable_add_dim[measurable]:
- "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
+ "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)"
(is "?f \<in> measurable ?P ?I")
proof (rule measurable_PiM_single)
fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)"
have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
- (if j = i then space (Pi\<^isub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
+ (if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
also have "\<dots> \<in> sets ?P"
using A j
@@ -469,11 +469,11 @@
qed (auto simp: space_pair_measure space_PiM PiE_def)
lemma measurable_component_update:
- "x \<in> space (Pi\<^isub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)"
+ "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
by simp
lemma measurable_merge[measurable]:
- "merge I J \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
+ "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)"
(is "?f \<in> measurable ?P ?U")
proof (rule measurable_PiM_single)
fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
@@ -488,7 +488,7 @@
lemma measurable_restrict[measurable (raw)]:
assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
- shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^isub>M I M)"
+ shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)"
proof (rule measurable_PiM_single)
fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N"
@@ -497,11 +497,11 @@
using A X by (auto intro!: measurable_sets)
qed (insert X, auto simp add: PiE_def dest: measurable_space)
-lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)"
+lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
by (intro measurable_restrict measurable_component_singleton) auto
lemma measurable_prod_emb[intro, simp]:
- "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^isub>M L M)"
+ "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
unfolding prod_emb_def space_PiM[symmetric]
by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
@@ -544,15 +544,15 @@
lemma (in finite_product_sigma_finite) sigma_finite_pairs:
"\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
(\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
- (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
- (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space (PiM I M)"
+ (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and>
+ (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)"
proof -
have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
using M.sigma_finite_incseq by metis
from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>"
by auto
- let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
+ let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k"
note space_PiM[simp]
show ?thesis
proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
@@ -573,14 +573,14 @@
qed
lemma
- shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\<lambda>k. undefined}"
- and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\<lambda>k. undefined} }"
+ shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
+ and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }"
by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
proof -
let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)"
- have "emeasure (Pi\<^isub>M {} M) (prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = 1"
+ have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
proof (subst emeasure_extend_measure_Pair[OF PiM_def])
show "positive (PiM {} M) ?\<mu>"
by (auto simp: positive_def)
@@ -588,7 +588,7 @@
by (rule sets.countably_additiveI_finite)
(auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: )
qed (auto simp: prod_emb_def)
- also have "(prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
+ also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
by (auto simp: prod_emb_def)
finally show ?thesis
by simp
@@ -598,7 +598,7 @@
by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def)
lemma (in product_sigma_finite) emeasure_PiM:
- "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+ "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
proof (induct I arbitrary: A rule: finite_induct)
case (insert i I)
interpret finite_product_sigma_finite M I by default fact
@@ -606,31 +606,31 @@
interpret I': finite_product_sigma_finite M "insert i I" by default fact
let ?h = "(\<lambda>(f, y). f(i := y))"
- let ?P = "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h"
+ let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h"
let ?\<mu> = "emeasure ?P"
let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))"
- have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>E (insert i I) A)) =
+ have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) =
(\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
proof (subst emeasure_extend_measure_Pair[OF PiM_def])
fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))"
then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto
- let ?p = "prod_emb (insert i I) M J (Pi\<^isub>E J E)"
- let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^isub>E j\<in>J-{i}. E j)"
+ let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)"
+ let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)"
have "?\<mu> ?p =
- emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i))"
+ emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))"
by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
- also have "?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
+ also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
using J E[rule_format, THEN sets.sets_into_space]
by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: split_if_asm)
- also have "emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
- emeasure (Pi\<^isub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
+ also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
+ emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto
- also have "?p' = (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
+ also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
using J E[rule_format, THEN sets.sets_into_space]
by (auto simp: prod_emb_iff PiE_def Pi_iff split: split_if_asm) blast+
- also have "emeasure (Pi\<^isub>M I M) (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
+ also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
(\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
using E by (subst insert) (auto intro!: setprod_cong)
also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
@@ -640,10 +640,10 @@
using insert(1,2) J E by (intro setprod_mono_one_right) auto
finally show "?\<mu> ?p = \<dots>" .
- show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \<in> Pow (\<Pi>\<^isub>E i\<in>insert i I. space (M i))"
+ show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))"
using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def)
next
- show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>"
+ show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>"
using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
next
show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
@@ -662,38 +662,38 @@
from sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
then have F: "\<And>j. j \<in> I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
- "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k)"
- "(\<Union>k. \<Pi>\<^isub>E j \<in> I. F j k) = space (Pi\<^isub>M I M)"
+ "incseq (\<lambda>k. \<Pi>\<^sub>E j \<in> I. F j k)"
+ "(\<Union>k. \<Pi>\<^sub>E j \<in> I. F j k) = space (Pi\<^sub>M I M)"
"\<And>k. \<And>j. j \<in> I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>"
by blast+
- let ?F = "\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k"
+ let ?F = "\<lambda>k. \<Pi>\<^sub>E j \<in> I. F j k"
show ?thesis
proof (unfold_locales, intro exI[of _ ?F] conjI allI)
- show "range ?F \<subseteq> sets (Pi\<^isub>M I M)" using F(1) `finite I` by auto
+ show "range ?F \<subseteq> sets (Pi\<^sub>M I M)" using F(1) `finite I` by auto
next
- from F(3) show "(\<Union>i. ?F i) = space (Pi\<^isub>M I M)" by simp
+ from F(3) show "(\<Union>i. ?F i) = space (Pi\<^sub>M I M)" by simp
next
fix j
from F `finite I` setprod_PInf[of I, OF emeasure_nonneg, of M]
- show "emeasure (\<Pi>\<^isub>M i\<in>I. M i) (?F j) \<noteq> \<infinity>"
+ show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (?F j) \<noteq> \<infinity>"
by (subst emeasure_PiM) auto
qed
qed
-sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^isub>M I M"
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
using sigma_finite[OF finite_index] .
lemma (in finite_product_sigma_finite) measure_times:
- "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^isub>M I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+ "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
using emeasure_PiM[OF finite_index] by auto
lemma (in product_sigma_finite) positive_integral_empty:
assumes pos: "0 \<le> f (\<lambda>k. undefined)"
- shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
+ shows "integral\<^sup>P (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
proof -
interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
- have "\<And>A. emeasure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
+ have "\<And>A. emeasure (Pi\<^sub>M {} M) (Pi\<^sub>E {} A) = 1"
using assms by (subst measure_times) auto
then show ?thesis
unfolding positive_integral_def simple_function_def simple_integral_def[abs_def]
@@ -707,38 +707,38 @@
lemma (in product_sigma_finite) distr_merge:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
- shows "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M) (merge I J) = Pi\<^isub>M (I \<union> J) M"
+ shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
(is "?D = ?P")
proof -
interpret I: finite_product_sigma_finite M I by default fact
interpret J: finite_product_sigma_finite M J by default fact
have "finite (I \<union> J)" using fin by auto
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
- interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
+ interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default
let ?g = "merge I J"
from IJ.sigma_finite_pairs obtain F where
F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
- "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
- "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space ?P"
+ "incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k)"
+ "(\<Union>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k) = space ?P"
"\<And>k. \<forall>i\<in>I\<union>J. emeasure (M i) (F i k) \<noteq> \<infinity>"
by auto
- let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
+ let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k"
show ?thesis
proof (rule measure_eqI_generator_eq[symmetric])
show "Int_stable (prod_algebra (I \<union> J) M)"
by (rule Int_stable_prod_algebra)
- show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^isub>E i \<in> I \<union> J. space (M i))"
+ show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^sub>E i \<in> I \<union> J. space (M i))"
by (rule prod_algebra_sets_into_space)
- show "sets ?P = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
+ show "sets ?P = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
by (rule sets_PiM)
- then show "sets ?D = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
+ then show "sets ?D = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
by simp
show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F
using fin by (auto simp: prod_algebra_eq_finite)
- show "(\<Union>i. \<Pi>\<^isub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i))"
+ show "(\<Union>i. \<Pi>\<^sub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i))"
using F(3) by (simp add: space_PiM)
next
fix k
@@ -746,13 +746,13 @@
show "emeasure ?P (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
next
fix A assume A: "A \<in> prod_algebra (I \<union> J) M"
- with fin obtain F where A_eq: "A = (Pi\<^isub>E (I \<union> J) F)" and F: "\<forall>i\<in>J. F i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
+ with fin obtain F where A_eq: "A = (Pi\<^sub>E (I \<union> J) F)" and F: "\<forall>i\<in>J. F i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
by (auto simp add: prod_algebra_eq_finite)
- let ?B = "Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M"
+ let ?B = "Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M"
let ?X = "?g -` A \<inter> space ?B"
- have "Pi\<^isub>E I F \<subseteq> space (Pi\<^isub>M I M)" "Pi\<^isub>E J F \<subseteq> space (Pi\<^isub>M J M)"
+ have "Pi\<^sub>E I F \<subseteq> space (Pi\<^sub>M I M)" "Pi\<^sub>E J F \<subseteq> space (Pi\<^sub>M J M)"
using F[rule_format, THEN sets.sets_into_space] by (force simp: space_PiM)+
- then have X: "?X = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
+ then have X: "?X = (Pi\<^sub>E I F \<times> Pi\<^sub>E J F)"
unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM)
have "emeasure ?D A = emeasure ?B ?X"
using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM)
@@ -761,7 +761,7 @@
by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times)
also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
using `finite J` `finite I` `I \<inter> J = {}` by (simp add: setprod_Un_one)
- also have "\<dots> = emeasure ?P (Pi\<^isub>E (I \<union> J) F)"
+ also have "\<dots> = emeasure ?P (Pi\<^sub>E (I \<union> J) F)"
using `finite J` `finite I` F unfolding A
by (intro IJ.measure_times[symmetric]) auto
finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp
@@ -770,14 +770,14 @@
lemma (in product_sigma_finite) product_positive_integral_fold:
assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
- and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
- shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
- (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
+ and f: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
+ shows "integral\<^sup>P (Pi\<^sub>M (I \<union> J) M) f =
+ (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
proof -
interpret I: finite_product_sigma_finite M I by default fact
interpret J: finite_product_sigma_finite M J by default fact
- interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
- have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+ interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default
+ have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
using measurable_comp[OF measurable_merge f] by (simp add: comp_def)
show ?thesis
apply (subst distr_merge[OF IJ, symmetric])
@@ -788,11 +788,11 @@
qed
lemma (in product_sigma_finite) distr_singleton:
- "distr (Pi\<^isub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
+ "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
proof (intro measure_eqI[symmetric])
interpret I: finite_product_sigma_finite M "{i}" by default simp
fix A assume A: "A \<in> sets (M i)"
- moreover then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M {i} M) = (\<Pi>\<^isub>E i\<in>{i}. A)"
+ moreover then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)"
using sets.sets_into_space by (auto simp: space_PiM)
ultimately show "emeasure (M i) A = emeasure ?D A"
using A I.measure_times[of "\<lambda>_. A"]
@@ -801,7 +801,7 @@
lemma (in product_sigma_finite) product_positive_integral_singleton:
assumes f: "f \<in> borel_measurable (M i)"
- shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
+ shows "integral\<^sup>P (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>P (M i) f"
proof -
interpret I: finite_product_sigma_finite M "{i}" by default simp
from f show ?thesis
@@ -813,8 +813,8 @@
lemma (in product_sigma_finite) product_positive_integral_insert:
assumes I[simp]: "finite I" "i \<notin> I"
- and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
- shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
+ and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
+ shows "integral\<^sup>P (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))"
proof -
interpret I: finite_product_sigma_finite M I by default auto
interpret i: finite_product_sigma_finite M "{i}" by default auto
@@ -823,12 +823,12 @@
show ?thesis
unfolding product_positive_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
proof (rule positive_integral_cong, subst product_positive_integral_singleton[symmetric])
- fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
+ fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
let ?f = "\<lambda>y. f (x(i := y))"
show "?f \<in> borel_measurable (M i)"
using measurable_comp[OF measurable_component_update f, OF x `i \<notin> I`]
unfolding comp_def .
- show "(\<integral>\<^isup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^isub>M {i} M) = (\<integral>\<^isup>+ y. f (x(i := y i)) \<partial>Pi\<^isub>M {i} M)"
+ show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)"
using x
by (auto intro!: positive_integral_cong arg_cong[where f=f]
simp add: space_PiM extensional_def PiE_def)
@@ -839,14 +839,14 @@
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
- shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
+ shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>P (M i) (f i))"
using assms proof induct
case (insert i I)
note `finite I`[intro, simp]
interpret I: finite_product_sigma_finite M I by default auto
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
using insert by (auto intro!: setprod_cong)
- have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
+ have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
using sets.sets_into_space insert
by (intro borel_measurable_ereal_setprod
measurable_comp[OF measurable_component_singleton, unfolded comp_def])
@@ -861,7 +861,7 @@
lemma (in product_sigma_finite) product_integral_singleton:
assumes f: "f \<in> borel_measurable (M i)"
- shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
+ shows "(\<integral>x. f (x i) \<partial>Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f"
proof -
interpret I: finite_product_sigma_finite M "{i}" by default simp
have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)"
@@ -871,7 +871,7 @@
unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
qed
lemma (in product_sigma_finite) distr_component:
- "distr (M i) (Pi\<^isub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^isub>M {i} M" (is "?D = ?P")
+ "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
proof (intro measure_eqI[symmetric])
interpret I: finite_product_sigma_finite M "{i}" by default simp
@@ -879,32 +879,32 @@
by (auto simp: extensional_def restrict_def)
fix A assume A: "A \<in> sets ?P"
- then have "emeasure ?P A = (\<integral>\<^isup>+x. indicator A x \<partial>?P)"
+ then have "emeasure ?P A = (\<integral>\<^sup>+x. indicator A x \<partial>?P)"
by simp
- also have "\<dots> = (\<integral>\<^isup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) (x i) \<partial>PiM {i} M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) (x i) \<partial>PiM {i} M)"
by (intro positive_integral_cong) (auto simp: space_PiM indicator_def PiE_def eq)
also have "\<dots> = emeasure ?D A"
using A by (simp add: product_positive_integral_singleton emeasure_distr)
- finally show "emeasure (Pi\<^isub>M {i} M) A = emeasure ?D A" .
+ finally show "emeasure (Pi\<^sub>M {i} M) A = emeasure ?D A" .
qed simp
lemma (in product_sigma_finite) product_integral_fold:
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
- and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
- shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
+ and f: "integrable (Pi\<^sub>M (I \<union> J) M) f"
+ shows "integral\<^sup>L (Pi\<^sub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^sub>M J M) \<partial>Pi\<^sub>M I M)"
proof -
interpret I: finite_product_sigma_finite M I by default fact
interpret J: finite_product_sigma_finite M J by default fact
have "finite (I \<union> J)" using fin by auto
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
- interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
+ interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default
let ?M = "merge I J"
let ?f = "\<lambda>x. f (?M x)"
- from f have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
+ from f have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
by auto
- have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+ have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def)
- have f_int: "integrable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ?f"
+ have f_int: "integrable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) ?f"
by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f)
show ?thesis
apply (subst distr_merge[symmetric, OF IJ fin])
@@ -915,16 +915,16 @@
qed
lemma (in product_sigma_finite)
- assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+ assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
shows emeasure_fold_integral:
- "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
+ "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I)
and emeasure_fold_measurable:
- "(\<lambda>x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
+ "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B)
proof -
interpret I: finite_product_sigma_finite M I by default fact
interpret J: finite_product_sigma_finite M J by default fact
- interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
- have merge: "merge I J -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+ interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
+ have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
by (intro measurable_sets[OF _ A] measurable_merge assms)
show ?I
@@ -941,22 +941,22 @@
lemma (in product_sigma_finite) product_integral_insert:
assumes I: "finite I" "i \<notin> I"
- and f: "integrable (Pi\<^isub>M (insert i I) M) f"
- shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
+ and f: "integrable (Pi\<^sub>M (insert i I) M) f"
+ shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
proof -
- have "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = integral\<^isup>L (Pi\<^isub>M (I \<union> {i}) M) f"
+ have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \<union> {i}) M) f"
by simp
- also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) \<partial>Pi\<^isub>M I M)"
+ also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)"
using f I by (intro product_integral_fold) auto
- also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
+ also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
proof (rule integral_cong, subst product_integral_singleton[symmetric])
- fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
- have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
+ fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
+ have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
using f by auto
show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`]
unfolding comp_def .
- from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^isub>M {i} M)"
+ from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^sub>M {i} M)"
by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def)
qed
finally show ?thesis .
@@ -965,28 +965,28 @@
lemma (in product_sigma_finite) product_integrable_setprod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
- shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
+ shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
proof -
interpret finite_product_sigma_finite M I by default fact
have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
using integrable unfolding integrable_def by auto
- have borel: "?f \<in> borel_measurable (Pi\<^isub>M I M)"
+ have borel: "?f \<in> borel_measurable (Pi\<^sub>M I M)"
using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto simp: comp_def)
- moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
+ moreover have "integrable (Pi\<^sub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
proof (unfold integrable_def, intro conjI)
- show "(\<lambda>x. abs (?f x)) \<in> borel_measurable (Pi\<^isub>M I M)"
+ show "(\<lambda>x. abs (?f x)) \<in> borel_measurable (Pi\<^sub>M I M)"
using borel by auto
- have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>Pi\<^isub>M I M) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>Pi\<^isub>M I M)"
+ have "(\<integral>\<^sup>+x. ereal (abs (?f x)) \<partial>Pi\<^sub>M I M) = (\<integral>\<^sup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>Pi\<^sub>M I M)"
by (simp add: setprod_ereal abs_setprod)
- also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))"
+ also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+x. ereal (abs (f i x)) \<partial>M i))"
using f by (subst product_positive_integral_setprod) auto
also have "\<dots> < \<infinity>"
using integrable[THEN integrable_abs]
by (simp add: setprod_PInf integrable_def positive_integral_positive)
- finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by auto
- have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) = (\<integral>\<^isup>+x. 0 \<partial>(Pi\<^isub>M I M))"
+ finally show "(\<integral>\<^sup>+x. ereal (abs (?f x)) \<partial>(Pi\<^sub>M I M)) \<noteq> \<infinity>" by auto
+ have "(\<integral>\<^sup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^sub>M I M)) = (\<integral>\<^sup>+x. 0 \<partial>(Pi\<^sub>M I M))"
by (intro positive_integral_cong_pos) auto
- then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by simp
+ then show "(\<integral>\<^sup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^sub>M I M)) \<noteq> \<infinity>" by simp
qed
ultimately show ?thesis
by (rule integrable_abs_iff[THEN iffD1])
@@ -995,17 +995,17 @@
lemma (in product_sigma_finite) product_integral_setprod:
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
- shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
+ shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>L (M i) (f i))"
using assms proof induct
case empty
- interpret finite_measure "Pi\<^isub>M {} M"
+ interpret finite_measure "Pi\<^sub>M {} M"
by rule (simp add: space_PiM)
show ?case by (simp add: space_PiM measure_def)
next
case (insert i I)
then have iI: "finite (insert i I)" by auto
then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
- integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
+ integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset)
interpret I: finite_product_sigma_finite M I by default fact
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
@@ -1016,7 +1016,7 @@
qed
lemma sets_Collect_single:
- "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^isub>M I M). x i \<in> A } \<in> sets (Pi\<^isub>M I M)"
+ "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
by simp
lemma sigma_prod_algebra_sigma_eq_infinite:
@@ -1025,16 +1025,16 @@
and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
- defines "P == {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> E i}"
+ defines "P == {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> E i}"
shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
proof
- let ?P = "sigma (space (Pi\<^isub>M I M)) P"
- have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
+ let ?P = "sigma (space (Pi\<^sub>M I M)) P"
+ have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>M I M))"
using E_closed by (auto simp: space_PiM P_def subset_eq)
- then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ then have space_P: "space ?P = (\<Pi>\<^sub>E i\<in>I. space (M i))"
by (simp add: space_PiM)
have "sets (PiM I M) =
- sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+ sigma_sets (space ?P) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
using sets_PiM_single[of I M] by (simp add: space_P)
also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
proof (safe intro!: sets.sigma_sets_subset)
@@ -1052,7 +1052,7 @@
from measurable_sets[OF this, of A] A `i \<in> I` E_closed
have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
by (simp add: E_generates)
- also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
+ also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A}"
using P_closed by (auto simp: space_PiM)
finally show "\<dots> \<in> sets ?P" .
qed
@@ -1070,19 +1070,19 @@
and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
- defines "P == { Pi\<^isub>E I F | F. \<forall>i\<in>I. F i \<in> E i }"
+ defines "P == { Pi\<^sub>E I F | F. \<forall>i\<in>I. F i \<in> E i }"
shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
proof
- let ?P = "sigma (space (Pi\<^isub>M I M)) P"
+ let ?P = "sigma (space (Pi\<^sub>M I M)) P"
from `finite I`[THEN ex_bij_betw_finite_nat] guess T ..
then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i"
by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f)
- have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
+ have P_closed: "P \<subseteq> Pow (space (Pi\<^sub>M I M))"
using E_closed by (auto simp: space_PiM P_def subset_eq)
- then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ then have space_P: "space ?P = (\<Pi>\<^sub>E i\<in>I. space (M i))"
by (simp add: space_PiM)
have "sets (PiM I M) =
- sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+ sigma_sets (space ?P) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
using sets_PiM_single[of I M] by (simp add: space_P)
also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
proof (safe intro!: sets.sigma_sets_subset)
@@ -1094,11 +1094,11 @@
show "\<forall>A\<in>E i. (\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
proof
fix A assume A: "A \<in> E i"
- then have "(\<lambda>x. x i) -` A \<inter> space ?P = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
+ then have "(\<lambda>x. x i) -` A \<inter> space ?P = (\<Pi>\<^sub>E j\<in>I. if i = j then A else space (M j))"
using E_closed `i \<in> I` by (auto simp: space_P subset_eq split: split_if_asm)
- also have "\<dots> = (\<Pi>\<^isub>E j\<in>I. \<Union>n. if i = j then A else S j n)"
+ also have "\<dots> = (\<Pi>\<^sub>E j\<in>I. \<Union>n. if i = j then A else S j n)"
by (intro PiE_cong) (simp add: S_union)
- also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j))"
+ also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>\<^sub>E j\<in>I. if i = j then A else S j (xs ! T j))"
using T
apply (auto simp: PiE_iff bchoice_iff)
apply (rule_tac x="map (\<lambda>n. f (the_inv_into I T n)) [0..<card I]" in exI)
@@ -1106,7 +1106,7 @@
done
also have "\<dots> \<in> sets ?P"
proof (safe intro!: sets.countable_UN)
- fix xs show "(\<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P"
+ fix xs show "(\<Pi>\<^sub>E j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P"
using A S_in_E
by (simp add: P_closed)
(auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"])
@@ -1118,7 +1118,7 @@
from measurable_sets[OF this, of A] A `i \<in> I` E_closed
have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
by (simp add: E_generates)
- also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
+ also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A}"
using P_closed by (auto simp: space_PiM)
finally show "\<dots> \<in> sets ?P" .
qed
@@ -1132,21 +1132,21 @@
lemma pair_measure_eq_distr_PiM:
fixes M1 :: "'a measure" and M2 :: "'a measure"
assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
- shows "(M1 \<Otimes>\<^isub>M M2) = distr (Pi\<^isub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. (x True, x False))"
+ shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
(is "?P = ?D")
proof (rule pair_measure_eqI[OF assms])
interpret B: product_sigma_finite "bool_case M1 M2"
unfolding product_sigma_finite_def using assms by (auto split: bool.split)
- let ?B = "Pi\<^isub>M UNIV (bool_case M1 M2)"
+ let ?B = "Pi\<^sub>M UNIV (bool_case M1 M2)"
have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
by auto
fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))"
by (simp add: UNIV_bool ac_simps)
- also have "\<dots> = emeasure ?B (Pi\<^isub>E UNIV (bool_case A B))"
+ also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (bool_case A B))"
using A B by (subst B.emeasure_PiM) (auto split: bool.split)
- also have "Pi\<^isub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
+ also have "Pi\<^sub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space]
by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split)
finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
--- a/src/HOL/Probability/Independent_Family.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Independent_Family.thy Tue Aug 13 16:25:47 2013 +0200
@@ -829,16 +829,16 @@
assumes "I \<noteq> {}"
assumes rv: "\<And>i. random_variable (M' i) (X i)"
shows "indep_vars M' X I \<longleftrightarrow>
- distr M (\<Pi>\<^isub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i))"
+ distr M (\<Pi>\<^sub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i))"
proof -
- let ?P = "\<Pi>\<^isub>M i\<in>I. M' i"
+ let ?P = "\<Pi>\<^sub>M i\<in>I. M' i"
let ?X = "\<lambda>x. \<lambda>i\<in>I. X i x"
let ?D = "distr M ?P ?X"
have X: "random_variable ?P ?X" by (intro measurable_restrict rv)
interpret D: prob_space ?D by (intro prob_space_distr X)
let ?D' = "\<lambda>i. distr M (M' i) (X i)"
- let ?P' = "\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i)"
+ let ?P' = "\<Pi>\<^sub>M i\<in>I. distr M (M' i) (X i)"
interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv)
interpret P: product_prob_space ?D' I ..
@@ -855,10 +855,10 @@
by (simp add: sets_PiM space_PiM)
show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')"
by (simp add: sets_PiM space_PiM cong: prod_algebra_cong)
- let ?A = "\<lambda>i. \<Pi>\<^isub>E i\<in>I. space (M' i)"
- show "range ?A \<subseteq> prod_algebra I M'" "(\<Union>i. ?A i) = space (Pi\<^isub>M I M')"
+ let ?A = "\<lambda>i. \<Pi>\<^sub>E i\<in>I. space (M' i)"
+ show "range ?A \<subseteq> prod_algebra I M'" "(\<Union>i. ?A i) = space (Pi\<^sub>M I M')"
by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong)
- { fix i show "emeasure ?D (\<Pi>\<^isub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto }
+ { fix i show "emeasure ?D (\<Pi>\<^sub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto }
next
fix E assume E: "E \<in> prod_algebra I M'"
from prod_algebraE[OF E] guess J Y . note J = this
@@ -896,7 +896,7 @@
qed
from bchoice[OF this] obtain Y where
Y: "\<And>j. j \<in> J \<Longrightarrow> Y' j = X j -` Y j \<inter> space M" "\<And>j. j \<in> J \<Longrightarrow> Y j \<in> sets (M' j)" by auto
- let ?E = "prod_emb I M' J (Pi\<^isub>E J Y)"
+ let ?E = "prod_emb I M' J (Pi\<^sub>E J Y)"
from Y have "(\<Inter>j\<in>J. Y' j) = ?X -` ?E \<inter> space M"
using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: split_if_asm)
then have "emeasure M (\<Inter>j\<in>J. Y' j) = emeasure M (?X -` ?E \<inter> space M)"
@@ -946,18 +946,18 @@
lemma (in prob_space) indep_var_distribution_eq:
"indep_var S X T Y \<longleftrightarrow> random_variable S X \<and> random_variable T Y \<and>
- distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" (is "_ \<longleftrightarrow> _ \<and> _ \<and> ?S \<Otimes>\<^isub>M ?T = ?J")
+ distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" (is "_ \<longleftrightarrow> _ \<and> _ \<and> ?S \<Otimes>\<^sub>M ?T = ?J")
proof safe
assume "indep_var S X T Y"
then show rvs: "random_variable S X" "random_variable T Y"
by (blast dest: indep_var_rv1 indep_var_rv2)+
- then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ then have XY: "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
by (rule measurable_Pair)
interpret X: prob_space ?S by (rule prob_space_distr) fact
interpret Y: prob_space ?T by (rule prob_space_distr) fact
interpret XY: pair_prob_space ?S ?T ..
- show "?S \<Otimes>\<^isub>M ?T = ?J"
+ show "?S \<Otimes>\<^sub>M ?T = ?J"
proof (rule pair_measure_eqI)
show "sigma_finite_measure ?S" ..
show "sigma_finite_measure ?T" ..
@@ -973,7 +973,7 @@
qed simp
next
assume rvs: "random_variable S X" "random_variable T Y"
- then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ then have XY: "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
by (rule measurable_Pair)
let ?S = "distr M S X" and ?T = "distr M T Y"
@@ -981,7 +981,7 @@
interpret Y: prob_space ?T by (rule prob_space_distr) fact
interpret XY: pair_prob_space ?S ?T ..
- assume "?S \<Otimes>\<^isub>M ?T = ?J"
+ assume "?S \<Otimes>\<^sub>M ?T = ?J"
{ fix S and X
have "Int_stable {X -` A \<inter> space M |A. A \<in> sets S}"
@@ -1004,8 +1004,8 @@
fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
then have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = emeasure ?J (A \<times> B)"
using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"])
- also have "\<dots> = emeasure (?S \<Otimes>\<^isub>M ?T) (A \<times> B)"
- unfolding `?S \<Otimes>\<^isub>M ?T = ?J` ..
+ also have "\<dots> = emeasure (?S \<Otimes>\<^sub>M ?T) (A \<times> B)"
+ unfolding `?S \<Otimes>\<^sub>M ?T = ?J` ..
also have "\<dots> = emeasure ?S A * emeasure ?T B"
using ab by (simp add: Y.emeasure_pair_measure_Times)
finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
@@ -1019,7 +1019,7 @@
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes X: "distributed M S X Px" and Y: "distributed M T Y Py"
assumes indep: "indep_var S X T Y"
- shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
+ shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
using indep_var_distribution_eq[of S X T Y] indep
by (intro distributed_joint_indep'[OF S T X Y]) auto
--- a/src/HOL/Probability/Infinite_Product_Measure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -10,11 +10,11 @@
lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
- shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
+ shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
proof cases
assume "finite I" with X show ?thesis by simp
next
- let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space (M i)"
+ let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space (M i)"
let ?G = generator
assume "\<not> finite I"
then have I_not_empty: "I \<noteq> {}" by auto
@@ -31,19 +31,19 @@
moreover from Z guess K' X' by (rule generatorE)
moreover def K \<equiv> "insert k K'"
moreover def X \<equiv> "emb K K' X'"
- ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
- "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
+ ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^sub>M K M)" "Z = emb I K X"
+ "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^sub>M K M) X"
by (auto simp: subset_insertI)
- let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
- { fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
+ let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M)"
+ { fix y assume y: "y \<in> space (Pi\<^sub>M J M)"
note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
moreover
- have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"
+ have **: "?M y \<in> sets (Pi\<^sub>M (K - J) M)"
using J K y by (intro merge_sets) auto
ultimately
- have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"
+ have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<in> ?G"
using J K by (intro generatorI) auto
- have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"
+ have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) = emeasure (Pi\<^sub>M (K - J) M) (?M y)"
unfolding * using K J by (subst mu_G_eq[OF _ _ _ **]) auto
note * ** *** this }
note merge_in_G = this
@@ -53,28 +53,28 @@
interpret J: finite_product_prob_space M J by default fact+
interpret KmJ: finite_product_prob_space M "K - J" by default fact+
- have "\<mu>G Z = emeasure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
+ have "\<mu>G Z = emeasure (Pi\<^sub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
using K J by simp
- also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. emeasure (Pi\<^sub>M (K - J) M) (?M x) \<partial>Pi\<^sub>M J M)"
using K J by (subst emeasure_fold_integral) auto
- also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
- (is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")
+ also have "\<dots> = (\<integral>\<^sup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<partial>Pi\<^sub>M J M)"
+ (is "_ = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)")
proof (intro positive_integral_cong)
- fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
+ fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
with K merge_in_G(2)[OF this]
- show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
+ show "emeasure (Pi\<^sub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst mu_G_eq) auto
qed
- finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" .
+ finally have fold: "\<mu>G Z = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)" .
- { fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
+ { fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
then have "\<mu>G (?MZ x) \<le> 1"
unfolding merge_in_G(4)[OF x] `Z = emb I K X`
by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
note le_1 = this
- let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^isub>M I M))"
- have "?q \<in> borel_measurable (Pi\<^isub>M J M)"
+ let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^sub>M I M))"
+ have "?q \<in> borel_measurable (Pi\<^sub>M J M)"
unfolding `Z = emb I K X` using J K merge_in_G(3)
by (simp add: merge_in_G mu_G_eq emeasure_fold_measurable cong: measurable_cong)
note this fold le_1 merge_in_G(3) }
@@ -98,14 +98,14 @@
using A positive_mu_G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
ultimately have "0 < ?a" by auto
- have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (limP J M (\<lambda>J. (Pi\<^isub>M J M))) X"
+ have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (limP J M (\<lambda>J. (Pi\<^sub>M J M))) X"
using A by (intro allI generator_Ex) auto
- then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)"
+ then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^sub>M (J' n) M)"
and A': "\<And>n. A n = emb I (J' n) (X' n)"
unfolding choice_iff by blast
moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
- ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^isub>M (J n) M)"
+ ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^sub>M (J n) M)"
by auto
with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"
unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)
@@ -119,7 +119,7 @@
using mu_G_spec[of "J 0" "A 0" "X 0"] J A_eq
by (auto intro!: INF_lower2[of 0] J.measure_le_1)
- let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)"
+ let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)"
{ fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
@@ -127,22 +127,22 @@
interpret J': finite_product_prob_space M J' by default fact+
let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)"
- let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^isub>M J' M)"
+ let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^sub>M J' M)"
{ fix n
- have "?q n \<in> borel_measurable (Pi\<^isub>M J' M)"
+ have "?q n \<in> borel_measurable (Pi\<^sub>M J' M)"
using Z J' by (intro fold(1)) auto
- then have "?Q n \<in> sets (Pi\<^isub>M J' M)"
+ then have "?Q n \<in> sets (Pi\<^sub>M J' M)"
by (rule measurable_sets) auto }
note Q_sets = this
- have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^isub>M J' M) (?Q n))"
+ have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^sub>M J' M) (?Q n))"
proof (intro INF_greatest)
fix n
have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)"
+ also have "\<dots> \<le> (\<integral>\<^sup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^sub>M J' M)"
unfolding fold(2)[OF J' `Z n \<in> ?G`]
proof (intro positive_integral_mono)
- fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
+ fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
then have "?q n x \<le> 1 + 0"
using J' Z fold(3) Z_sets by auto
also have "\<dots> \<le> 1 + ?a / 2^(k+1)"
@@ -151,21 +151,21 @@
with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
by (auto split: split_indicator simp del: power_Suc)
qed
- also have "\<dots> = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"
+ also have "\<dots> = emeasure (Pi\<^sub>M J' M) (?Q n) + ?a / 2^(k+1)"
using `0 \<le> ?a` Q_sets J'.emeasure_space_1
by (subst positive_integral_add) auto
- finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
- by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"])
+ finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^sub>M J' M) (?Q n)" using `?a \<le> 1`
+ by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^sub>M J' M) (?Q n)"])
(auto simp: field_simps)
qed
- also have "\<dots> = emeasure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
+ also have "\<dots> = emeasure (Pi\<^sub>M J' M) (\<Inter>n. ?Q n)"
proof (intro INF_emeasure_decseq)
- show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto
+ show "range ?Q \<subseteq> sets (Pi\<^sub>M J' M)" using Q_sets by auto
show "decseq ?Q"
unfolding decseq_def
proof (safe intro!: vimageI[OF refl])
fix m n :: nat assume "m \<le> n"
- fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
+ fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
assume "?a / 2^(k+1) \<le> ?q n x"
also have "?q n x \<le> ?q m x"
proof (rule mu_G_mono)
@@ -179,7 +179,7 @@
qed simp
finally have "(\<Inter>n. ?Q n) \<noteq> {}"
using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
- then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
+ then have "\<exists>w\<in>space (Pi\<^sub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
note Ex_w = this
let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)"
@@ -188,18 +188,18 @@
from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this
let ?P =
- "\<lambda>k wk w. w \<in> space (Pi\<^isub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and>
+ "\<lambda>k wk w. w \<in> space (Pi\<^sub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and>
(\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)"
def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))"
- { fix k have w: "w k \<in> space (Pi\<^isub>M (J k) M) \<and>
+ { fix k have w: "w k \<in> space (Pi\<^sub>M (J k) M) \<and>
(\<forall>n. ?a / 2 ^ (k + 1) \<le> ?q k n (w k)) \<and> (k \<noteq> 0 \<longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1))"
proof (induct k)
case 0 with w0 show ?case
unfolding w_def nat_rec_0 by auto
next
case (Suc k)
- then have wk: "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
+ then have wk: "w k \<in> space (Pi\<^sub>M (J k) M)" by auto
have "\<exists>w'. ?P k (w k) w'"
proof cases
assume [simp]: "J k = J (Suc k)"
@@ -211,7 +211,7 @@
also have "\<dots> \<le> ?q k n (w k)" using Suc by auto
finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp
next
- show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)"
+ show "w k \<in> space (Pi\<^sub>M (J (Suc k)) M)"
using Suc by simp
then show "restrict (w k) (J k) = w k"
by (simp add: extensional_restrict space_PiM)
@@ -225,7 +225,7 @@
using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc
by (auto simp: decseq_def)
from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
- obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)"
+ obtain w' where w': "w' \<in> space (Pi\<^sub>M ?D M)"
"\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto
let ?w = "merge (J k) ?D (w k, w')"
have [simp]: "\<And>x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) =
@@ -246,7 +246,7 @@
then show ?case by auto
qed
moreover
- then have "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
+ then have "w k \<in> space (Pi\<^sub>M (J k) M)" by auto
moreover
from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto
then have "?M (J k) (A k) (w k) \<noteq> {}"
@@ -255,9 +255,9 @@
then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto
then have "\<exists>x\<in>A k. restrict x (J k) = w k"
- using `w k \<in> space (Pi\<^isub>M (J k) M)`
+ using `w k \<in> space (Pi\<^sub>M (J k) M)`
by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM)
- ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)"
+ ultimately have "w k \<in> space (Pi\<^sub>M (J k) M)"
"\<exists>x\<in>A k. restrict x (J k) = w k"
"k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)"
by auto }
@@ -291,7 +291,7 @@
using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }
note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
- have w': "w' \<in> space (Pi\<^isub>M I M)"
+ have w': "w' \<in> space (Pi\<^sub>M I M)"
using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)
{ fix n
@@ -312,40 +312,40 @@
by (simp add: Pi_iff)
next
fix J X assume J: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
- then show "emb I J (Pi\<^isub>E J X) \<in> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ then show "emb I J (Pi\<^sub>E J X) \<in> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
by (auto simp: Pi_iff prod_emb_def dest: sets.sets_into_space)
- have "emb I J (Pi\<^isub>E J X) \<in> generator"
+ have "emb I J (Pi\<^sub>E J X) \<in> generator"
using J `I \<noteq> {}` by (intro generatorI') (auto simp: Pi_iff)
- then have "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))"
+ then have "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))"
using \<mu> by simp
also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
using J `I \<noteq> {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
- finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = \<dots>" .
+ finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = \<dots>" .
next
let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"
have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)"
using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
- emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
+ emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
using X by (auto simp add: emeasure_PiM)
next
- show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>"
+ show "positive (sets (Pi\<^sub>M I M)) \<mu>" "countably_additive (sets (Pi\<^sub>M I M)) \<mu>"
using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def)
qed
qed
-sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^isub>M I M"
+sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^sub>M I M"
proof
- show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1"
+ show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1"
proof cases
assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
next
assume "I \<noteq> {}"
then obtain i where "i \<in> I" by auto
- moreover then have "emb I {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i)) = (space (Pi\<^isub>M I M))"
+ moreover then have "emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)) = (space (Pi\<^sub>M I M))"
by (auto simp: prod_emb_def space_PiM)
ultimately show ?thesis
using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
@@ -355,10 +355,10 @@
lemma (in product_prob_space) emeasure_PiM_emb:
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
- shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
+ shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
proof cases
assume "J = {}"
- moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^isub>M I M)"
+ moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^sub>M I M)"
by (auto simp: space_PiM prod_emb_def)
ultimately show ?thesis
by (simp add: space_PiM_empty P.emeasure_space_1)
@@ -369,22 +369,22 @@
lemma (in product_prob_space) emeasure_PiM_Collect:
assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
- shows "emeasure (Pi\<^isub>M I M) {x\<in>space (Pi\<^isub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))"
+ shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))"
proof -
- have "{x\<in>space (Pi\<^isub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^isub>E J X)"
+ have "{x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^sub>E J X)"
unfolding prod_emb_def using assms by (auto simp: space_PiM Pi_iff)
with emeasure_PiM_emb[OF assms] show ?thesis by simp
qed
lemma (in product_prob_space) emeasure_PiM_Collect_single:
assumes X: "i \<in> I" "A \<in> sets (M i)"
- shows "emeasure (Pi\<^isub>M I M) {x\<in>space (Pi\<^isub>M I M). x i \<in> A} = emeasure (M i) A"
+ shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). x i \<in> A} = emeasure (M i) A"
using emeasure_PiM_Collect[of "{i}" "\<lambda>i. A"] assms
by simp
lemma (in product_prob_space) measure_PiM_emb:
assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
- shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
+ shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
using emeasure_PiM_emb[OF assms]
unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)
@@ -394,7 +394,7 @@
by (simp add: space_PiM PiE_iff cong: conj_cong)
lemma (in finite_product_prob_space) finite_measure_PiM_emb:
- "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
+ "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M]
by auto
@@ -413,25 +413,25 @@
assumes "I \<noteq> {}"
assumes "sets M' = sets (PiM I M)"
assumes eq: "\<And>J F. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>j. j \<in> J \<Longrightarrow> F j \<in> sets (M j)) \<Longrightarrow>
- emeasure M' (prod_emb I M J (\<Pi>\<^isub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))"
+ emeasure M' (prod_emb I M J (\<Pi>\<^sub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))"
shows "M' = (PiM I M)"
proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space])
- show "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"
+ show "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
by (rule sets_PiM)
- then show "sets M' = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"
+ then show "sets M' = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
unfolding `sets M' = sets (PiM I M)` by simp
def i \<equiv> "SOME i. i \<in> I"
with `I \<noteq> {}` have i: "i \<in> I"
by (auto intro: someI_ex)
- def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^isub>E j\<in>{i}. space (M i))"
+ def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. space (M i))"
then show "range A \<subseteq> prod_algebra I M"
by (auto intro!: prod_algebraI i)
have A_eq: "\<And>i. A i = space (PiM I M)"
by (auto simp: prod_emb_def space_PiM Pi_iff A_def i)
- show "(\<Union>i. A i) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+ show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
unfolding A_eq by (auto simp: space_PiM)
show "\<And>i. emeasure (PiM I M) (A i) \<noteq> \<infinity>"
unfolding A_eq P.emeasure_space_1 by simp
@@ -459,22 +459,22 @@
by (auto simp: comb_seq_def)
lemma measurable_comb_seq:
- "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) (\<Pi>\<^isub>M i\<in>UNIV. M)"
+ "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) (\<Pi>\<^sub>M i\<in>UNIV. M)"
proof (rule measurable_PiM_single)
- show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^isub>E space M)"
+ show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^sub>E space M)"
by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq)
fix j :: nat and A assume A: "A \<in> sets M"
- then have *: "{\<omega> \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} =
- (if j < i then {\<omega> \<in> space (\<Pi>\<^isub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^isub>M i\<in>UNIV. M)
- else space (\<Pi>\<^isub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^isub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
+ then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} =
+ (if j < i then {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^sub>M i\<in>UNIV. M)
+ else space (\<Pi>\<^sub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space)
- show "{\<omega> \<in> space ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^isub>M i\<in>UNIV. M) \<Otimes>\<^isub>M (\<Pi>\<^isub>M i\<in>UNIV. M))"
+ show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). prod_case (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))"
unfolding * by (auto simp: A intro!: sets_Collect_single)
qed
lemma measurable_comb_seq'[measurable (raw)]:
- assumes f: "f \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
- shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
+ assumes f: "f \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
+ shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
using measurable_compose[OF measurable_Pair[OF f g] measurable_comb_seq] by simp
lemma comb_seq_0: "comb_seq 0 \<omega> \<omega>' = \<omega>'"
@@ -502,13 +502,13 @@
locale sequence_space = product_prob_space "\<lambda>i. M" "UNIV :: nat set" for M
begin
-abbreviation "S \<equiv> \<Pi>\<^isub>M i\<in>UNIV::nat set. M"
+abbreviation "S \<equiv> \<Pi>\<^sub>M i\<in>UNIV::nat set. M"
lemma infprod_in_sets[intro]:
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"
shows "Pi UNIV E \<in> sets S"
proof -
- have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^isub>E j\<in>{..i}. E j))"
+ have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^sub>E j\<in>{..i}. E j))"
using E E[THEN sets.sets_into_space]
by (auto simp: prod_emb_def Pi_iff extensional_def) blast
with E show ?thesis by auto
@@ -518,7 +518,7 @@
fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"
shows "(\<lambda>n. \<Prod>i\<le>n. measure M (E i)) ----> measure S (Pi UNIV E)"
proof -
- let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)"
+ let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^sub>E {.. n} E)"
have "\<And>n. (\<Prod>i\<le>n. measure M (E i)) = measure S (?E n)"
using E by (simp add: measure_PiM_emb)
moreover have "Pi UNIV E = (\<Inter>n. ?E n)"
@@ -538,21 +538,21 @@
by auto
lemma PiM_comb_seq:
- "distr (S \<Otimes>\<^isub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _")
+ "distr (S \<Otimes>\<^sub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _")
proof (rule PiM_eq)
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"
let "distr _ _ ?f" = "?D"
fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M"
- let ?X = "prod_emb ?I ?M J (\<Pi>\<^isub>E j\<in>J. E j)"
+ let ?X = "prod_emb ?I ?M J (\<Pi>\<^sub>E j\<in>J. E j)"
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"
using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)
- with J have "?f -` ?X \<inter> space (S \<Otimes>\<^isub>M S) =
+ with J have "?f -` ?X \<inter> space (S \<Otimes>\<^sub>M S) =
(prod_emb ?I ?M (J \<inter> {..<i}) (PIE j:J \<inter> {..<i}. E j)) \<times>
(prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \<times> ?F")
by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib PiE_iff
split: split_comb_seq split_comb_seq_asm)
- then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^isub>M S) (?E \<times> ?F)"
+ then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^sub>M S) (?E \<times> ?F)"
by (subst emeasure_distr[OF measurable_comb_seq])
(auto intro!: sets_PiM_I simp: split_beta' J)
also have "\<dots> = emeasure S ?E * emeasure S ?F"
@@ -570,7 +570,7 @@
qed simp_all
lemma PiM_iter:
- "distr (M \<Otimes>\<^isub>M S) S (\<lambda>(s, \<omega>). nat_case s \<omega>) = S" (is "?D = _")
+ "distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). nat_case s \<omega>) = S" (is "?D = _")
proof (rule PiM_eq)
let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"
let "distr _ _ ?f" = "?D"
@@ -579,11 +579,11 @@
let ?X = "prod_emb ?I ?M J (PIE j:J. E j)"
have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"
using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)
- with J have "?f -` ?X \<inter> space (M \<Otimes>\<^isub>M S) = (if 0 \<in> J then E 0 else space M) \<times>
+ with J have "?f -` ?X \<inter> space (M \<Otimes>\<^sub>M S) = (if 0 \<in> J then E 0 else space M) \<times>
(prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \<times> ?F")
by (auto simp: space_pair_measure space_PiM PiE_iff prod_emb_def all_conj_distrib
split: nat.split nat.split_asm)
- then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^isub>M S) (?E \<times> ?F)"
+ then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^sub>M S) (?E \<times> ?F)"
by (subst emeasure_distr)
(auto intro!: sets_PiM_I simp: split_beta' J)
also have "\<dots> = emeasure M ?E * emeasure S ?F"
--- a/src/HOL/Probability/Information.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Information.thy Tue Aug 13 16:25:47 2013 +0200
@@ -77,7 +77,7 @@
"entropy_density b M N = log b \<circ> real \<circ> RN_deriv M N"
definition
- "KL_divergence b M N = integral\<^isup>L N (entropy_density b M N)"
+ "KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)"
lemma (in information_space) measurable_entropy_density:
assumes ac: "absolutely_continuous M N" "sets N = events"
@@ -153,21 +153,21 @@
have [simp, intro]: "?D_set \<in> sets M"
using D by auto
- have D_neg: "(\<integral>\<^isup>+ x. ereal (- D x) \<partial>M) = 0"
+ have D_neg: "(\<integral>\<^sup>+ x. ereal (- D x) \<partial>M) = 0"
using D by (subst positive_integral_0_iff_AE) auto
- have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = emeasure (density M D) (space M)"
+ have "(\<integral>\<^sup>+ x. ereal (D x) \<partial>M) = emeasure (density M D) (space M)"
using D by (simp add: emeasure_density cong: positive_integral_cong)
- then have D_pos: "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = 1"
+ then have D_pos: "(\<integral>\<^sup>+ x. ereal (D x) \<partial>M) = 1"
using N.emeasure_space_1 by simp
- have "integrable M D" "integral\<^isup>L M D = 1"
+ have "integrable M D" "integral\<^sup>L M D = 1"
using D D_pos D_neg unfolding integrable_def lebesgue_integral_def by simp_all
have "0 \<le> 1 - measure M ?D_set"
using prob_le_1 by (auto simp: field_simps)
also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)"
- using `integrable M D` `integral\<^isup>L M D = 1`
+ using `integrable M D` `integral\<^sup>L M D = 1`
by (simp add: emeasure_eq_measure)
also have "\<dots> < (\<integral> x. D x * (ln b * log b (D x)) \<partial>M)"
proof (rule integral_less_AE)
@@ -185,13 +185,13 @@
then have disj: "AE x in M. D x = 1 \<or> D x = 0"
using D(1) by (auto intro!: AE_I[OF subset_refl] sets.sets_Collect)
- have "emeasure M {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
+ have "emeasure M {x\<in>space M. D x = 1} = (\<integral>\<^sup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
using D(1) by auto
- also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. ereal (D x) \<partial>M)"
using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def)
finally have "AE x in M. D x = 1"
using D D_pos by (intro AE_I_eq_1) auto
- then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. ereal (D x) * indicator A x\<partial>M)"
+ then have "(\<integral>\<^sup>+x. indicator A x\<partial>M) = (\<integral>\<^sup>+x. ereal (D x) * indicator A x\<partial>M)"
by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
also have "\<dots> = density M D A"
using `A \<in> sets M` D by (simp add: emeasure_density)
@@ -261,7 +261,7 @@
then have "AE x in M. log b (real (RN_deriv M M x)) = 0"
by (elim AE_mp) simp
from integral_cong_AE[OF this]
- have "integral\<^isup>L M (entropy_density b M M) = 0"
+ have "integral\<^sup>L M (entropy_density b M M) = 0"
by (simp add: entropy_density_def comp_def)
then show "KL_divergence b M M = 0"
unfolding KL_divergence_def
@@ -378,13 +378,13 @@
lemma ac_fst:
assumes "sigma_finite_measure T"
- shows "absolutely_continuous S (distr (S \<Otimes>\<^isub>M T) S fst)"
+ shows "absolutely_continuous S (distr (S \<Otimes>\<^sub>M T) S fst)"
proof -
interpret sigma_finite_measure T by fact
{ fix A assume "A \<in> sets S" "emeasure S A = 0"
- moreover then have "fst -` A \<inter> space (S \<Otimes>\<^isub>M T) = A \<times> space T"
+ moreover then have "fst -` A \<inter> space (S \<Otimes>\<^sub>M T) = A \<times> space T"
by (auto simp: space_pair_measure dest!: sets.sets_into_space)
- ultimately have "emeasure (S \<Otimes>\<^isub>M T) (fst -` A \<inter> space (S \<Otimes>\<^isub>M T)) = 0"
+ ultimately have "emeasure (S \<Otimes>\<^sub>M T) (fst -` A \<inter> space (S \<Otimes>\<^sub>M T)) = 0"
by (simp add: emeasure_pair_measure_Times) }
then show ?thesis
unfolding absolutely_continuous_def
@@ -395,13 +395,13 @@
lemma ac_snd:
assumes "sigma_finite_measure T"
- shows "absolutely_continuous T (distr (S \<Otimes>\<^isub>M T) T snd)"
+ shows "absolutely_continuous T (distr (S \<Otimes>\<^sub>M T) T snd)"
proof -
interpret sigma_finite_measure T by fact
{ fix A assume "A \<in> sets T" "emeasure T A = 0"
- moreover then have "snd -` A \<inter> space (S \<Otimes>\<^isub>M T) = space S \<times> A"
+ moreover then have "snd -` A \<inter> space (S \<Otimes>\<^sub>M T) = space S \<times> A"
by (auto simp: space_pair_measure dest!: sets.sets_into_space)
- ultimately have "emeasure (S \<Otimes>\<^isub>M T) (snd -` A \<inter> space (S \<Otimes>\<^isub>M T)) = 0"
+ ultimately have "emeasure (S \<Otimes>\<^sub>M T) (snd -` A \<inter> space (S \<Otimes>\<^sub>M T)) = 0"
by (simp add: emeasure_pair_measure_Times) }
then show ?thesis
unfolding absolutely_continuous_def
@@ -456,12 +456,12 @@
definition (in prob_space)
"mutual_information b S T X Y =
- KL_divergence b (distr M S X \<Otimes>\<^isub>M distr M T Y) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)))"
+ KL_divergence b (distr M S X \<Otimes>\<^sub>M distr M T Y) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))"
lemma (in information_space) mutual_information_indep_vars:
fixes S T X Y
- defines "P \<equiv> distr M S X \<Otimes>\<^isub>M distr M T Y"
- defines "Q \<equiv> distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ defines "P \<equiv> distr M S X \<Otimes>\<^sub>M distr M T Y"
+ defines "Q \<equiv> distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
shows "indep_var S X T Y \<longleftrightarrow>
(random_variable S X \<and> random_variable T Y \<and>
absolutely_continuous P Q \<and> integrable Q (entropy_density b P Q) \<and>
@@ -480,7 +480,7 @@
interpret Q: prob_space Q unfolding Q_def
by (rule prob_space_distr) simp
- { assume "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ { assume "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
then have [simp]: "Q = P" unfolding Q_def P_def by simp
show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def)
@@ -515,7 +515,7 @@
show "KL_divergence b P Q = 0"
using I_eq_0 unfolding mutual_information_def by (simp add: P_def Q_def)
qed
- then show "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ then show "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
unfolding P_def Q_def .. }
qed
@@ -527,16 +527,16 @@
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Fx: "finite_entropy S X Px" and Fy: "finite_entropy T Y Py"
- assumes Fxy: "finite_entropy (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Fxy: "finite_entropy (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
- shows mutual_information_distr': "mutual_information b S T X Y = integral\<^isup>L (S \<Otimes>\<^isub>M T) f" (is "?M = ?R")
+ shows mutual_information_distr': "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R")
and mutual_information_nonneg': "0 \<le> mutual_information b S T X Y"
proof -
have Px: "distributed M S X Px"
using Fx by (auto simp: finite_entropy_def)
have Py: "distributed M T Y Py"
using Fy by (auto simp: finite_entropy_def)
- have Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ have Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
using Fxy by (auto simp: finite_entropy_def)
have X: "random_variable S X"
@@ -549,7 +549,7 @@
interpret X: prob_space "distr M S X" using X by (rule prob_space_distr)
interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr)
interpret XY: pair_prob_space "distr M S X" "distr M T Y" ..
- let ?P = "S \<Otimes>\<^isub>M T"
+ let ?P = "S \<Otimes>\<^sub>M T"
let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
{ fix A assume "A \<in> sets S"
@@ -566,7 +566,7 @@
have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
by auto
- have distr_eq: "distr M S X \<Otimes>\<^isub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
+ have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
proof (subst pair_measure_density)
show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
@@ -604,16 +604,16 @@
have X: "X = fst \<circ> (\<lambda>x. (X x, Y x))" and Y: "Y = snd \<circ> (\<lambda>x. (X x, Y x))"
by auto
- have "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))"
+ have "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))"
using finite_entropy_integrable[OF Fxy]
using finite_entropy_integrable_transform[OF Fx Pxy, of fst]
using finite_entropy_integrable_transform[OF Fy Pxy, of snd]
by simp
- moreover have "f \<in> borel_measurable (S \<Otimes>\<^isub>M T)"
+ moreover have "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)"
unfolding f_def using Px Py Pxy
by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd''
intro!: borel_measurable_times borel_measurable_log borel_measurable_divide)
- ultimately have int: "integrable (S \<Otimes>\<^isub>M T) f"
+ ultimately have int: "integrable (S \<Otimes>\<^sub>M T) f"
apply (rule integrable_cong_AE_imp)
using
distributed_transform_AE[OF measurable_fst ac_fst, of T, OF T Px]
@@ -627,10 +627,10 @@
show "0 \<le> ?M" unfolding M
proof (rule ST.KL_density_density_nonneg
[OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
- show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x))) "
+ show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x))) "
unfolding distributed_distr_eq_density[OF Pxy, symmetric]
using distributed_measurable[OF Pxy] by (rule prob_space_distr)
- show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+ show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
unfolding distr_eq[symmetric] by unfold_locales
qed
qed
@@ -640,10 +640,10 @@
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
assumes "sigma_finite_measure S" "sigma_finite_measure T"
assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
- shows mutual_information_distr: "mutual_information b S T X Y = integral\<^isup>L (S \<Otimes>\<^isub>M T) f" (is "?M = ?R")
- and mutual_information_nonneg: "integrable (S \<Otimes>\<^isub>M T) f \<Longrightarrow> 0 \<le> mutual_information b S T X Y"
+ shows mutual_information_distr: "mutual_information b S T X Y = integral\<^sup>L (S \<Otimes>\<^sub>M T) f" (is "?M = ?R")
+ and mutual_information_nonneg: "integrable (S \<Otimes>\<^sub>M T) f \<Longrightarrow> 0 \<le> mutual_information b S T X Y"
proof -
have X: "random_variable S X"
using Px by (auto simp: distributed_def)
@@ -655,7 +655,7 @@
interpret X: prob_space "distr M S X" using X by (rule prob_space_distr)
interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr)
interpret XY: pair_prob_space "distr M S X" "distr M T Y" ..
- let ?P = "S \<Otimes>\<^isub>M T"
+ let ?P = "S \<Otimes>\<^sub>M T"
let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
{ fix A assume "A \<in> sets S"
@@ -672,7 +672,7 @@
have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
by auto
- have distr_eq: "distr M S X \<Otimes>\<^isub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
+ have distr_eq: "distr M S X \<Otimes>\<^sub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
proof (subst pair_measure_density)
show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
@@ -707,14 +707,14 @@
using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac
by (rule ST.KL_density_density)
- assume int: "integrable (S \<Otimes>\<^isub>M T) f"
+ assume int: "integrable (S \<Otimes>\<^sub>M T) f"
show "0 \<le> ?M" unfolding M
proof (rule ST.KL_density_density_nonneg
[OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
- show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x))) "
+ show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x))) "
unfolding distributed_distr_eq_density[OF Pxy, symmetric]
using distributed_measurable[OF Pxy] by (rule prob_space_distr)
- show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+ show "prob_space (density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
unfolding distr_eq[symmetric] by unfold_locales
qed
qed
@@ -723,7 +723,7 @@
fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
assumes "sigma_finite_measure S" "sigma_finite_measure T"
assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y"
shows mutual_information_eq_0: "mutual_information b S T X Y = 0"
proof -
@@ -731,18 +731,18 @@
interpret T: sigma_finite_measure T by fact
interpret ST: pair_sigma_finite S T ..
- have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
+ have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
moreover
- have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+ have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
moreover
- have "AE x in S \<Otimes>\<^isub>M T. Pxy x = Px (fst x) * Py (snd x)"
+ have "AE x in S \<Otimes>\<^sub>M T. Pxy x = Px (fst x) * Py (snd x)"
using distributed_real_measurable[OF Px] distributed_real_measurable[OF Py] distributed_real_measurable[OF Pxy]
by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'')
- ultimately have "AE x in S \<Otimes>\<^isub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0"
+ ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0"
by eventually_elim simp
- then have "(\<integral>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
+ then have "(\<integral>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \<partial>(S \<Otimes>\<^sub>M T)) = (\<integral>x. 0 \<partial>(S \<Otimes>\<^sub>M T))"
by (rule integral_cong_AE)
then show ?thesis
by (subst mutual_information_distr[OF assms(1-5)]) simp
@@ -762,7 +762,7 @@
let ?f = "\<lambda>x. ?Pxy x * log b (?Pxy x / (Px (fst x) * Py (snd x)))"
have "\<And>x. ?f x = (if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) else 0)"
by auto
- with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M))) =
+ with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M))) =
(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum_cases split_beta'
intro!: setsum_cong)
@@ -852,7 +852,7 @@
with Px have "AE x in MX. Px x = 0"
by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ereal_iff)
moreover
- from X.emeasure_space_1 have "(\<integral>\<^isup>+x. Px x \<partial>MX) = 1"
+ from X.emeasure_space_1 have "(\<integral>\<^sup>+x. Px x \<partial>MX) = 1"
unfolding distributed_distr_eq_density[OF X] using Px
by (subst (asm) emeasure_density)
(auto simp: borel_measurable_ereal_iff intro!: integral_cong cong: positive_integral_cong)
@@ -886,7 +886,7 @@
using Px by (auto simp: AE_density)
have [simp]: "\<And>x. x \<in> space MX \<Longrightarrow> ereal (if Px x = 0 then 0 else 1) = indicator {x \<in> space MX. Px x \<noteq> 0} x"
by (auto simp: one_ereal_def)
- have "(\<integral>\<^isup>+ x. max 0 (ereal (- (if Px x = 0 then 0 else 1))) \<partial>MX) = (\<integral>\<^isup>+ x. 0 \<partial>MX)"
+ have "(\<integral>\<^sup>+ x. max 0 (ereal (- (if Px x = 0 then 0 else 1))) \<partial>MX) = (\<integral>\<^sup>+ x. 0 \<partial>MX)"
by (intro positive_integral_cong) (auto split: split_max)
then show "integrable (distr M MX X) (\<lambda>x. 1 / Px x)"
unfolding distributed_distr_eq_density[OF X] using Px
@@ -983,7 +983,7 @@
definition (in prob_space)
"conditional_mutual_information b MX MY MZ X Y Z \<equiv>
- mutual_information b MX (MY \<Otimes>\<^isub>M MZ) X (\<lambda>x. (Y x, Z x)) -
+ mutual_information b MX (MY \<Otimes>\<^sub>M MZ) X (\<lambda>x. (Y x, Z x)) -
mutual_information b MX MZ X Z"
abbreviation (in information_space)
@@ -995,13 +995,13 @@
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
assumes Px[measurable]: "distributed M S X Px"
assumes Pz[measurable]: "distributed M P Z Pz"
- assumes Pyz[measurable]: "distributed M (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
- assumes Pxz[measurable]: "distributed M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) Pxz"
- assumes Pxyz[measurable]: "distributed M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
- assumes I1: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
- assumes I2: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
+ assumes Pyz[measurable]: "distributed M (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
+ assumes Pxz[measurable]: "distributed M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) Pxz"
+ assumes Pxyz[measurable]: "distributed M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+ assumes I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
+ assumes I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
shows conditional_mutual_information_generic_eq: "conditional_mutual_information b S T P X Y Z
- = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" (is "?eq")
+ = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" (is "?eq")
and conditional_mutual_information_generic_nonneg: "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg")
proof -
interpret S: sigma_finite_measure S by fact
@@ -1010,47 +1010,47 @@
interpret TP: pair_sigma_finite T P ..
interpret SP: pair_sigma_finite S P ..
interpret ST: pair_sigma_finite S T ..
- interpret SPT: pair_sigma_finite "S \<Otimes>\<^isub>M P" T ..
- interpret STP: pair_sigma_finite S "T \<Otimes>\<^isub>M P" ..
- interpret TPS: pair_sigma_finite "T \<Otimes>\<^isub>M P" S ..
- have TP: "sigma_finite_measure (T \<Otimes>\<^isub>M P)" ..
- have SP: "sigma_finite_measure (S \<Otimes>\<^isub>M P)" ..
- have YZ: "random_variable (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x))"
+ interpret SPT: pair_sigma_finite "S \<Otimes>\<^sub>M P" T ..
+ interpret STP: pair_sigma_finite S "T \<Otimes>\<^sub>M P" ..
+ interpret TPS: pair_sigma_finite "T \<Otimes>\<^sub>M P" S ..
+ have TP: "sigma_finite_measure (T \<Otimes>\<^sub>M P)" ..
+ have SP: "sigma_finite_measure (S \<Otimes>\<^sub>M P)" ..
+ have YZ: "random_variable (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x))"
using Pyz by (simp add: distributed_measurable)
- from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) =
- distr (distr M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). (x, z))"
+ from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) =
+ distr (distr M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). (x, z))"
by (simp add: comp_def distr_distr)
have "mutual_information b S P X Z =
- (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^isub>M P))"
+ (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))"
by (rule mutual_information_distr[OF S P Px Pz Pxz])
- also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+ also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
using b_gt_1 Pxz Px Pz
by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"]) (auto simp: split_beta')
finally have mi_eq:
- "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" .
+ "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" .
- have ae1: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
+ have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
- moreover have ae2: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) auto
- moreover have ae3: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) auto
- moreover have ae4: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) auto
- moreover have ae5: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
+ moreover have ae5: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Px (fst x)"
using Px by (intro STP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
- moreover have ae6: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
+ moreover have ae6: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pyz (snd x)"
using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
- moreover have ae7: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
+ moreover have ae7: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd (snd x))"
using Pz Pz[THEN distributed_real_measurable]
by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
- moreover have ae8: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
+ moreover have ae8: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure)
moreover note Pxyz[THEN distributed_real_AE]
- ultimately have ae: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
+ ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P.
Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
@@ -1073,12 +1073,12 @@
apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
done
- let ?P = "density (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) Pxyz"
+ let ?P = "density (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) Pxyz"
interpret P: prob_space ?P
unfolding distributed_distr_eq_density[OF Pxyz, symmetric]
by (rule prob_space_distr) simp
- let ?Q = "density (T \<Otimes>\<^isub>M P) Pyz"
+ let ?Q = "density (T \<Otimes>\<^sub>M P) Pyz"
interpret Q: prob_space ?Q
unfolding distributed_distr_eq_density[OF Pyz, symmetric]
by (rule prob_space_distr) simp
@@ -1086,15 +1086,15 @@
let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
from subdensity_real[of snd, OF _ Pyz Pz]
- have aeX1: "AE x in T \<Otimes>\<^isub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
- have aeX2: "AE x in T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd x)"
+ have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
+ have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)"
using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def dest: distributed_real_AE)
- have aeX3: "AE y in T \<Otimes>\<^isub>M P. (\<integral>\<^isup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+ have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
- have "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+ have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
apply (subst positive_integral_density)
apply simp
apply (rule distributed_AE[OF Pxyz])
@@ -1104,27 +1104,27 @@
apply eventually_elim
apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg)
done
- also have "\<dots> = (\<integral>\<^isup>+(y, z). \<integral>\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^isub>M P)"
+ also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)"
by (subst STP.positive_integral_snd_measurable[symmetric]) (auto simp add: split_beta')
- also have "\<dots> = (\<integral>\<^isup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^isub>M P)"
+ also have "\<dots> = (\<integral>\<^sup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
apply (rule positive_integral_cong_AE)
using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
apply eventually_elim
proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
- "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)"
- then show "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
+ "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)"
+ then show "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
by (subst positive_integral_multc)
(auto intro!: divide_nonneg_nonneg split: prod.split)
qed
also have "\<dots> = 1"
using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
by (subst positive_integral_density[symmetric]) auto
- finally have le1: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> 1" .
+ finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" .
also have "\<dots> < \<infinity>" by simp
- finally have fin: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
+ finally have fin: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
- have pos: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> 0"
+ have pos: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> 0"
apply (subst positive_integral_density)
apply simp
apply (rule distributed_AE[OF Pxyz])
@@ -1132,19 +1132,19 @@
apply (simp add: split_beta')
proof
let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
- assume "(\<integral>\<^isup>+ x. ?g x \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)) = 0"
- then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. ?g x \<le> 0"
+ assume "(\<integral>\<^sup>+ x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0"
+ then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x \<le> 0"
by (intro positive_integral_0_iff_AE[THEN iffD1]) auto
- then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxyz x = 0"
+ then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff)
- then have "(\<integral>\<^isup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) = 0"
+ then have "(\<integral>\<^sup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
by (subst positive_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
with P.emeasure_space_1 show False
by (subst (asm) emeasure_density) (auto cong: positive_integral_cong)
qed
- have neg: "(\<integral>\<^isup>+ x. - ?f x \<partial>?P) = 0"
+ have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0"
apply (rule positive_integral_0_iff_AE[THEN iffD2])
apply simp
apply (subst AE_density)
@@ -1154,22 +1154,22 @@
apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg)
done
- have I3: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+ have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]])
using ae
apply (auto simp: split_beta')
done
- have "- log b 1 \<le> - log b (integral\<^isup>L ?P ?f)"
+ have "- log b 1 \<le> - log b (integral\<^sup>L ?P ?f)"
proof (intro le_imp_neg_le log_le[OF b_gt_1])
- show "0 < integral\<^isup>L ?P ?f"
+ show "0 < integral\<^sup>L ?P ?f"
using neg pos fin positive_integral_positive[of ?P ?f]
- by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
- show "integral\<^isup>L ?P ?f \<le> 1"
+ by (cases "(\<integral>\<^sup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
+ show "integral\<^sup>L ?P ?f \<le> 1"
using neg le1 fin positive_integral_positive[of ?P ?f]
- by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
+ by (cases "(\<integral>\<^sup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
qed
- also have "- log b (integral\<^isup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
+ also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
show "AE x in ?P. ?f x \<in> {0<..}"
unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
@@ -1211,11 +1211,11 @@
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
assumes Fx: "finite_entropy S X Px"
assumes Fz: "finite_entropy P Z Pz"
- assumes Fyz: "finite_entropy (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
- assumes Fxz: "finite_entropy (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) Pxz"
- assumes Fxyz: "finite_entropy (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+ assumes Fyz: "finite_entropy (T \<Otimes>\<^sub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
+ assumes Fxz: "finite_entropy (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) Pxz"
+ assumes Fxyz: "finite_entropy (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
shows conditional_mutual_information_generic_eq': "conditional_mutual_information b S T P X Y Z
- = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" (is "?eq")
+ = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" (is "?eq")
and conditional_mutual_information_generic_nonneg': "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg")
proof -
note Px = Fx[THEN finite_entropy_distributed, measurable]
@@ -1230,45 +1230,45 @@
interpret TP: pair_sigma_finite T P ..
interpret SP: pair_sigma_finite S P ..
interpret ST: pair_sigma_finite S T ..
- interpret SPT: pair_sigma_finite "S \<Otimes>\<^isub>M P" T ..
- interpret STP: pair_sigma_finite S "T \<Otimes>\<^isub>M P" ..
- interpret TPS: pair_sigma_finite "T \<Otimes>\<^isub>M P" S ..
- have TP: "sigma_finite_measure (T \<Otimes>\<^isub>M P)" ..
- have SP: "sigma_finite_measure (S \<Otimes>\<^isub>M P)" ..
+ interpret SPT: pair_sigma_finite "S \<Otimes>\<^sub>M P" T ..
+ interpret STP: pair_sigma_finite S "T \<Otimes>\<^sub>M P" ..
+ interpret TPS: pair_sigma_finite "T \<Otimes>\<^sub>M P" S ..
+ have TP: "sigma_finite_measure (T \<Otimes>\<^sub>M P)" ..
+ have SP: "sigma_finite_measure (S \<Otimes>\<^sub>M P)" ..
- from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) =
- distr (distr M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). (x, z))"
+ from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Z x)) =
+ distr (distr M (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). (x, z))"
by (simp add: distr_distr comp_def)
have "mutual_information b S P X Z =
- (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^isub>M P))"
+ (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^sub>M P))"
by (rule mutual_information_distr[OF S P Px Pz Pxz])
- also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+ also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
using b_gt_1 Pxz Px Pz
by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"])
(auto simp: split_beta')
finally have mi_eq:
- "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" .
+ "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))" .
- have ae1: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
+ have ae1: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
- moreover have ae2: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae2: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) auto
- moreover have ae3: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae3: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) auto
- moreover have ae4: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+ moreover have ae4: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) auto
- moreover have ae5: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
+ moreover have ae5: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Px (fst x)"
using Px by (intro STP.AE_pair_measure) (auto dest: distributed_real_AE)
- moreover have ae6: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
+ moreover have ae6: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pyz (snd x)"
using Pyz by (intro STP.AE_pair_measure) (auto dest: distributed_real_AE)
- moreover have ae7: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
+ moreover have ae7: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd (snd x))"
using Pz Pz[THEN distributed_real_measurable] by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
- moreover have ae8: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
+ moreover have ae8: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def)
moreover note ae9 = Pxyz[THEN distributed_real_AE]
- ultimately have ae: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
+ ultimately have ae: "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P.
Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
@@ -1284,30 +1284,30 @@
qed simp
qed
- have "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)
+ have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)
(\<lambda>x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))"
using finite_entropy_integrable[OF Fxyz]
using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
using finite_entropy_integrable_transform[OF Fyz Pxyz, of snd]
by simp
- moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \<in> borel_measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)"
+ moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)"
using Pxyz Px Pyz by simp
- ultimately have I1: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
+ ultimately have I1: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
apply (rule integrable_cong_AE_imp)
using ae1 ae4 ae5 ae6 ae9
by eventually_elim
(auto simp: log_divide_eq log_mult_eq mult_nonneg_nonneg field_simps zero_less_mult_iff)
- have "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)
+ have "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)
(\<lambda>x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))"
using finite_entropy_integrable_transform[OF Fxz Pxyz, of "\<lambda>x. (fst x, snd (snd x))"]
using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
using finite_entropy_integrable_transform[OF Fz Pxyz, of "snd \<circ> snd"]
by simp
- moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \<in> borel_measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)"
+ moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \<in> borel_measurable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)"
using Pxyz Px Pz
by auto
- ultimately have I2: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
+ ultimately have I2: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
apply (rule integrable_cong_AE_imp)
using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 ae9
by eventually_elim
@@ -1321,25 +1321,25 @@
apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
done
- let ?P = "density (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) Pxyz"
+ let ?P = "density (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) Pxyz"
interpret P: prob_space ?P
unfolding distributed_distr_eq_density[OF Pxyz, symmetric] by (rule prob_space_distr) simp
- let ?Q = "density (T \<Otimes>\<^isub>M P) Pyz"
+ let ?Q = "density (T \<Otimes>\<^sub>M P) Pyz"
interpret Q: prob_space ?Q
unfolding distributed_distr_eq_density[OF Pyz, symmetric] by (rule prob_space_distr) simp
let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
from subdensity_real[of snd, OF _ Pyz Pz]
- have aeX1: "AE x in T \<Otimes>\<^isub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
- have aeX2: "AE x in T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd x)"
+ have aeX1: "AE x in T \<Otimes>\<^sub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
+ have aeX2: "AE x in T \<Otimes>\<^sub>M P. 0 \<le> Pz (snd x)"
using Pz by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
- have aeX3: "AE y in T \<Otimes>\<^isub>M P. (\<integral>\<^isup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+ have aeX3: "AE y in T \<Otimes>\<^sub>M P. (\<integral>\<^sup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
by (intro TP.AE_pair_measure) (auto dest: distributed_real_AE)
- have "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+ have "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P))"
apply (subst positive_integral_density)
apply (rule distributed_borel_measurable[OF Pxyz])
apply (rule distributed_AE[OF Pxyz])
@@ -1349,27 +1349,27 @@
apply eventually_elim
apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg)
done
- also have "\<dots> = (\<integral>\<^isup>+(y, z). \<integral>\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^isub>M P)"
+ also have "\<dots> = (\<integral>\<^sup>+(y, z). \<integral>\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^sub>M P)"
by (subst STP.positive_integral_snd_measurable[symmetric])
(auto simp add: split_beta')
- also have "\<dots> = (\<integral>\<^isup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^isub>M P)"
+ also have "\<dots> = (\<integral>\<^sup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^sub>M P)"
apply (rule positive_integral_cong_AE)
using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
apply eventually_elim
proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
- "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)"
- then show "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
+ "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)"
+ then show "(\<integral>\<^sup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
by (subst positive_integral_multc) (auto intro!: divide_nonneg_nonneg)
qed
also have "\<dots> = 1"
using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
by (subst positive_integral_density[symmetric]) auto
- finally have le1: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> 1" .
+ finally have le1: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<le> 1" .
also have "\<dots> < \<infinity>" by simp
- finally have fin: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
+ finally have fin: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
- have pos: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> 0"
+ have pos: "(\<integral>\<^sup>+ x. ?f x \<partial>?P) \<noteq> 0"
apply (subst positive_integral_density)
apply simp
apply (rule distributed_AE[OF Pxyz])
@@ -1377,19 +1377,19 @@
apply (simp add: split_beta')
proof
let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
- assume "(\<integral>\<^isup>+ x. ?g x \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)) = 0"
- then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. ?g x \<le> 0"
+ assume "(\<integral>\<^sup>+ x. ?g x \<partial>(S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P)) = 0"
+ then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. ?g x \<le> 0"
by (intro positive_integral_0_iff_AE[THEN iffD1]) (auto intro!: borel_measurable_ereal measurable_If)
- then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxyz x = 0"
+ then have "AE x in S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P. Pxyz x = 0"
using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff)
- then have "(\<integral>\<^isup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) = 0"
+ then have "(\<integral>\<^sup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) = 0"
by (subst positive_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
with P.emeasure_space_1 show False
by (subst (asm) emeasure_density) (auto cong: positive_integral_cong)
qed
- have neg: "(\<integral>\<^isup>+ x. - ?f x \<partial>?P) = 0"
+ have neg: "(\<integral>\<^sup>+ x. - ?f x \<partial>?P) = 0"
apply (rule positive_integral_0_iff_AE[THEN iffD2])
apply (auto simp: split_beta') []
apply (subst AE_density)
@@ -1399,22 +1399,22 @@
apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg)
done
- have I3: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+ have I3: "integrable (S \<Otimes>\<^sub>M T \<Otimes>\<^sub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]])
using ae
apply (auto simp: split_beta')
done
- have "- log b 1 \<le> - log b (integral\<^isup>L ?P ?f)"
+ have "- log b 1 \<le> - log b (integral\<^sup>L ?P ?f)"
proof (intro le_imp_neg_le log_le[OF b_gt_1])
- show "0 < integral\<^isup>L ?P ?f"
+ show "0 < integral\<^sup>L ?P ?f"
using neg pos fin positive_integral_positive[of ?P ?f]
- by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
- show "integral\<^isup>L ?P ?f \<le> 1"
+ by (cases "(\<integral>\<^sup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
+ show "integral\<^sup>L ?P ?f \<le> 1"
using neg le1 fin positive_integral_positive[of ?P ?f]
- by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
+ by (cases "(\<integral>\<^sup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
qed
- also have "- log b (integral\<^isup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
+ also have "- log b (integral\<^sup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
show "AE x in ?P. ?f x \<in> {0<..}"
unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
@@ -1468,13 +1468,13 @@
by (simp add: sigma_finite_measure_count_space_finite)
show "sigma_finite_measure (count_space (Z ` space M))"
by (simp add: sigma_finite_measure_count_space_finite)
- have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) \<Otimes>\<^isub>M count_space (Z ` space M) =
+ have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) \<Otimes>\<^sub>M count_space (Z ` space M) =
count_space (X`space M \<times> Y`space M \<times> Z`space M)"
(is "?P = ?C")
by (simp add: pair_measure_count_space)
let ?Px = "\<lambda>x. measure M (X -` {x} \<inter> space M)"
- have "(\<lambda>x. (X x, Z x)) \<in> measurable M (count_space (X ` space M) \<Otimes>\<^isub>M count_space (Z ` space M))"
+ have "(\<lambda>x. (X x, Z x)) \<in> measurable M (count_space (X ` space M) \<Otimes>\<^sub>M count_space (Z ` space M))"
using simple_distributed_joint[OF Pxz] by (rule distributed_measurable)
from measurable_comp[OF this measurable_fst]
have "random_variable (count_space (X ` space M)) X"
@@ -1505,7 +1505,7 @@
assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z"
shows "0 \<le> \<I>(X ; Y | Z)"
proof -
- have [simp]: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) \<Otimes>\<^isub>M count_space (Z ` space M) =
+ have [simp]: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) \<Otimes>\<^sub>M count_space (Z ` space M) =
count_space (X`space M \<times> Y`space M \<times> Z`space M)"
by (simp add: pair_measure_count_space X Y Z simple_functionD)
note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)]
@@ -1525,8 +1525,8 @@
subsection {* Conditional Entropy *}
definition (in prob_space)
- "conditional_entropy b S T X Y = - (\<integral>(x, y). log b (real (RN_deriv (S \<Otimes>\<^isub>M T) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (x, y)) /
- real (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)))"
+ "conditional_entropy b S T X Y = - (\<integral>(x, y). log b (real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (x, y)) /
+ real (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))"
abbreviation (in information_space)
conditional_entropy_Pow ("\<H>'(_ | _')") where
@@ -1536,20 +1536,20 @@
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Py[measurable]: "distributed M T Y Py"
- assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^isub>M T))"
+ assumes Pxy[measurable]: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^sub>M T))"
proof -
interpret S: sigma_finite_measure S by fact
interpret T: sigma_finite_measure T by fact
interpret ST: pair_sigma_finite S T ..
- have "AE x in density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x)). Pxy x = real (RN_deriv (S \<Otimes>\<^isub>M T) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) x)"
+ have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x)). Pxy x = real (RN_deriv (S \<Otimes>\<^sub>M T) (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) x)"
unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
unfolding distributed_distr_eq_density[OF Pxy]
using distributed_RN_deriv[OF Pxy]
by auto
moreover
- have "AE x in density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x)). Py (snd x) = real (RN_deriv T (distr M T Y) (snd x))"
+ have "AE x in density (S \<Otimes>\<^sub>M T) (\<lambda>x. ereal (Pxy x)). Py (snd x) = real (RN_deriv T (distr M T Y) (snd x))"
unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
unfolding distributed_distr_eq_density[OF Py]
apply (rule ST.AE_pair_measure)
@@ -1560,7 +1560,7 @@
apply auto
done
ultimately
- have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+ have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
unfolding conditional_entropy_def neg_equal_iff_equal
apply (subst integral_density(1)[symmetric])
apply (auto simp: distributed_real_measurable[OF Pxy] distributed_real_AE[OF Pxy]
@@ -1575,10 +1575,10 @@
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- assumes I1: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
- assumes I2: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
- shows "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes I1: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+ assumes I2: "integrable (S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+ shows "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
proof -
interpret S: sigma_finite_measure S by fact
interpret T: sigma_finite_measure T by fact
@@ -1586,31 +1586,31 @@
have "entropy b T Y = - (\<integral>y. Py y * log b (Py y) \<partial>T)"
by (rule entropy_distr[OF Py])
- also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))"
+ also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))"
using b_gt_1 Py[THEN distributed_real_measurable]
by (subst distributed_transform_integral[OF Pxy Py, where T=snd]) (auto intro!: integral_cong)
- finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))" .
+ finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^sub>M T))" .
- have ae2: "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+ have ae2: "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
- moreover have ae4: "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
+ moreover have ae4: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)"
using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
moreover note ae5 = Pxy[THEN distributed_real_AE]
- ultimately have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and>
+ ultimately have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and>
(Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Py (snd x)))"
by eventually_elim auto
- then have ae: "AE x in S \<Otimes>\<^isub>M T.
+ then have ae: "AE x in S \<Otimes>\<^sub>M T.
Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))"
by eventually_elim (auto simp: log_simps mult_pos_pos field_simps b_gt_1)
have "conditional_entropy b S T X Y =
- - (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+ - (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
unfolding conditional_entropy_generic_eq[OF S T Py Pxy] neg_equal_iff_equal
apply (intro integral_cong_AE)
using ae
apply eventually_elim
apply auto
done
- also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^isub>M T)) - - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+ also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) - - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
by (simp add: integral_diff[OF I1 I2])
finally show ?thesis
unfolding conditional_entropy_generic_eq[OF S T Py Pxy] entropy_distr[OF Pxy] e_eq
@@ -1619,9 +1619,9 @@
lemma (in information_space) conditional_entropy_eq_entropy_simple:
assumes X: "simple_function M X" and Y: "simple_function M Y"
- shows "\<H>(X | Y) = entropy b (count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x)) - \<H>(Y)"
+ shows "\<H>(X | Y) = entropy b (count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x)) - \<H>(Y)"
proof -
- have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
+ have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
(is "?P = ?C") using X Y by (simp add: simple_functionD pair_measure_count_space)
show ?thesis
by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite
@@ -1647,7 +1647,7 @@
show "sigma_finite_measure (count_space (Y ` space M))"
by (simp add: sigma_finite_measure_count_space_finite)
let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)"
- have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
+ have "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
(is "?P = ?C")
using Y by (simp add: simple_distributed_finite pair_measure_count_space)
have eq: "(\<lambda>(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) =
@@ -1704,20 +1704,20 @@
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- assumes Ix: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
- assumes Iy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
- assumes Ixy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
- shows "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
+ assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+ assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+ shows "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
proof -
- have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^isub>M T))"
+ have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^sub>M T))"
using b_gt_1 Px[THEN distributed_real_measurable]
apply (subst entropy_distr[OF Px])
apply (subst distributed_transform_integral[OF Pxy Px, where T=fst])
apply (auto intro!: integral_cong)
done
- have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+ have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^sub>M T))"
using b_gt_1 Py[THEN distributed_real_measurable]
apply (subst entropy_distr[OF Py])
apply (subst distributed_transform_integral[OF Pxy Py, where T=snd])
@@ -1727,21 +1727,21 @@
interpret S: sigma_finite_measure S by fact
interpret T: sigma_finite_measure T by fact
interpret ST: pair_sigma_finite S T ..
- have ST: "sigma_finite_measure (S \<Otimes>\<^isub>M T)" ..
+ have ST: "sigma_finite_measure (S \<Otimes>\<^sub>M T)" ..
- have XY: "entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^isub>M T))"
+ have XY: "entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^sub>M T))"
by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong)
- have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
+ have "AE x in S \<Otimes>\<^sub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
- moreover have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+ moreover have "AE x in S \<Otimes>\<^sub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
- moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Px (fst x)"
+ moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Px (fst x)"
using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
- moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
+ moreover have "AE x in S \<Otimes>\<^sub>M T. 0 \<le> Py (snd x)"
using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
moreover note Pxy[THEN distributed_real_AE]
- ultimately have "AE x in S \<Otimes>\<^isub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) =
+ ultimately have "AE x in S \<Otimes>\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) =
Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
(is "AE x in _. ?f x = ?g x")
proof eventually_elim
@@ -1756,7 +1756,7 @@
qed simp
qed
- have "entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = integral\<^isup>L (S \<Otimes>\<^isub>M T) ?f"
+ have "entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?f"
unfolding X Y XY
apply (subst integral_diff)
apply (intro integral_diff Ixy Ix Iy)+
@@ -1764,7 +1764,7 @@
apply (intro integral_diff Ixy Ix Iy)+
apply (simp add: field_simps)
done
- also have "\<dots> = integral\<^isup>L (S \<Otimes>\<^isub>M T) ?g"
+ also have "\<dots> = integral\<^sup>L (S \<Otimes>\<^sub>M T) ?g"
using `AE x in _. ?f x = ?g x` by (rule integral_cong_AE)
also have "\<dots> = mutual_information b S T X Y"
by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric])
@@ -1775,10 +1775,10 @@
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- assumes Ix: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
- assumes Iy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
- assumes Ixy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Ix: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
+ assumes Iy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+ assumes Ixy: "integrable(S \<Otimes>\<^sub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
shows "mutual_information b S T X Y = entropy b S X - conditional_entropy b S T X Y"
using
mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy]
@@ -1798,10 +1798,10 @@
then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))"
by (rule simple_distributedI) auto
from simple_distributed_joint_finite[OF this, simp]
- have eq: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
+ have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
by (simp add: pair_measure_count_space)
- have "\<I>(X ; Y) = \<H>(X) + \<H>(Y) - entropy b (count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x))"
+ have "\<I>(X ; Y) = \<H>(X) + \<H>(Y) - entropy b (count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x))"
using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]
by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space)
then show ?thesis
@@ -1823,7 +1823,7 @@
by (rule simple_distributedI) auto
from simple_distributed_joint_finite[OF this, simp]
- have eq: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
+ have eq: "count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
by (simp add: pair_measure_count_space)
show ?thesis
@@ -1844,7 +1844,7 @@
fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes Px: "finite_entropy S X Px" and Py: "finite_entropy T Y Py"
- assumes Pxy: "finite_entropy (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ assumes Pxy: "finite_entropy (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
shows "conditional_entropy b S T X Y \<le> entropy b S X"
proof -
@@ -1875,11 +1875,11 @@
by (subst (2) setsum_reindex) (auto simp: inj_on_def intro!: setsum_cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))"
by (auto intro!: setsum_cong)
- also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
+ also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
by (subst entropy_distr[OF simple_distributed_joint[OF YX]])
(auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
cong del: setsum_cong intro!: setsum_mono_zero_left)
- finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
+ finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
then show ?thesis
unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp
qed
--- a/src/HOL/Probability/Lebesgue_Integration.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Tue Aug 13 16:25:47 2013 +0200
@@ -491,18 +491,18 @@
section "Simple integral"
-definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^isup>S") where
- "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
+definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>S") where
+ "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
syntax
- "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
+ "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
translations
- "\<integral>\<^isup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
+ "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
lemma simple_integral_cong:
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "integral\<^isup>S M f = integral\<^isup>S M g"
+ shows "integral\<^sup>S M f = integral\<^sup>S M g"
proof -
have "f ` space M = g ` space M"
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
@@ -511,7 +511,7 @@
qed
lemma simple_integral_const[simp]:
- "(\<integral>\<^isup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
+ "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
proof (cases "space M = {}")
case True thus ?thesis unfolding simple_integral_def by simp
next
@@ -521,7 +521,7 @@
lemma simple_function_partition:
assumes f: "simple_function M f" and g: "simple_function M g"
- shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * (emeasure M) A)"
+ shows "integral\<^sup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * (emeasure M) A)"
(is "_ = setsum _ (?p ` space M)")
proof-
let ?sub = "\<lambda>x. ?p ` (f -` {x} \<inter> space M)"
@@ -545,7 +545,7 @@
have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
with sets have "(emeasure M) (f -` {f x} \<inter> space M) = setsum (emeasure M) (?sub (f x))"
by (subst setsum_emeasure) (auto simp: disjoint_family_on_def) }
- hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * (emeasure M) A)"
+ hence "integral\<^sup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * (emeasure M) A)"
unfolding simple_integral_def using f sets
by (subst setsum_Sigma[symmetric])
(auto intro!: setsum_cong setsum_ereal_right_distrib)
@@ -569,7 +569,7 @@
lemma simple_integral_add[simp]:
assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
- shows "(\<integral>\<^isup>Sx. f x + g x \<partial>M) = integral\<^isup>S M f + integral\<^isup>S M g"
+ shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g"
proof -
{ fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
assume "x \<in> space M"
@@ -588,7 +588,7 @@
lemma simple_integral_setsum[simp]:
assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
- shows "(\<integral>\<^isup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>S M (f i))"
+ shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))"
proof cases
assume "finite P"
from this assms show ?thesis
@@ -597,7 +597,7 @@
lemma simple_integral_mult[simp]:
assumes f: "simple_function M f" "\<And>x. 0 \<le> f x"
- shows "(\<integral>\<^isup>Sx. c * f x \<partial>M) = c * integral\<^isup>S M f"
+ shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f"
proof -
note mult = simple_function_mult[OF simple_function_const[of _ c] f(1)]
{ fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
@@ -614,7 +614,7 @@
lemma simple_integral_mono_AE:
assumes f: "simple_function M f" and g: "simple_function M g"
and mono: "AE x in M. f x \<le> g x"
- shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
+ shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
proof -
let ?S = "\<lambda>x. (g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)"
have *: "\<And>x. g -` {g x} \<inter> f -` {f x} \<inter> space M = ?S x"
@@ -651,19 +651,19 @@
lemma simple_integral_mono:
assumes "simple_function M f" and "simple_function M g"
and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
- shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
+ shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
using assms by (intro simple_integral_mono_AE) auto
lemma simple_integral_cong_AE:
assumes "simple_function M f" and "simple_function M g"
and "AE x in M. f x = g x"
- shows "integral\<^isup>S M f = integral\<^isup>S M g"
+ shows "integral\<^sup>S M f = integral\<^sup>S M g"
using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
lemma simple_integral_cong':
assumes sf: "simple_function M f" "simple_function M g"
and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
- shows "integral\<^isup>S M f = integral\<^isup>S M g"
+ shows "integral\<^sup>S M f = integral\<^sup>S M g"
proof (intro simple_integral_cong_AE sf AE_I)
show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
@@ -673,11 +673,11 @@
lemma simple_integral_indicator:
assumes "A \<in> sets M"
assumes f: "simple_function M f"
- shows "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
+ shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>x \<in> f ` space M. x * (emeasure M) (f -` {x} \<inter> space M \<inter> A))"
proof cases
assume "A = space M"
- moreover hence "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) = integral\<^isup>S M f"
+ moreover hence "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = integral\<^sup>S M f"
by (auto intro!: simple_integral_cong)
moreover have "\<And>X. X \<inter> space M \<inter> space M = X \<inter> space M" by auto
ultimately show ?thesis by (simp add: simple_integral_def)
@@ -693,7 +693,7 @@
next
show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x])
qed
- have *: "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
+ have *: "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
(\<Sum>x \<in> f ` space M \<union> {0}. x * (emeasure M) (f -` {x} \<inter> space M \<inter> A))"
unfolding simple_integral_def I
proof (rule setsum_mono_zero_cong_left)
@@ -719,7 +719,7 @@
lemma simple_integral_indicator_only[simp]:
assumes "A \<in> sets M"
- shows "integral\<^isup>S M (indicator A) = emeasure M A"
+ shows "integral\<^sup>S M (indicator A) = emeasure M A"
proof cases
assume "space M = {}" hence "A = {}" using sets.sets_into_space[OF assms] by auto
thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
@@ -733,55 +733,55 @@
lemma simple_integral_null_set:
assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
- shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
+ shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
proof -
have "AE x in M. indicator N x = (0 :: ereal)"
using `N \<in> null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N])
- then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
+ then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)"
using assms apply (intro simple_integral_cong_AE) by auto
then show ?thesis by simp
qed
lemma simple_integral_cong_AE_mult_indicator:
assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M"
- shows "integral\<^isup>S M f = (\<integral>\<^isup>Sx. f x * indicator S x \<partial>M)"
+ shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)"
using assms by (intro simple_integral_cong_AE) auto
lemma simple_integral_cmult_indicator:
assumes A: "A \<in> sets M"
- shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * (emeasure M) A"
+ shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * (emeasure M) A"
using simple_integral_mult[OF simple_function_indicator[OF A]]
unfolding simple_integral_indicator_only[OF A] by simp
lemma simple_integral_positive:
assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x"
- shows "0 \<le> integral\<^isup>S M f"
+ shows "0 \<le> integral\<^sup>S M f"
proof -
- have "integral\<^isup>S M (\<lambda>x. 0) \<le> integral\<^isup>S M f"
+ have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f"
using simple_integral_mono_AE[OF _ f ae] by auto
then show ?thesis by simp
qed
section "Continuous positive integration"
-definition positive_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^isup>P") where
- "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
+definition positive_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>P") where
+ "integral\<^sup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^sup>S M g)"
syntax
- "_positive_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
+ "_positive_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>+ _. _ \<partial>_" [60,61] 110)
translations
- "\<integral>\<^isup>+ x. f \<partial>M" == "CONST positive_integral M (%x. f)"
+ "\<integral>\<^sup>+ x. f \<partial>M" == "CONST positive_integral M (%x. f)"
lemma positive_integral_positive:
- "0 \<le> integral\<^isup>P M f"
+ "0 \<le> integral\<^sup>P M f"
by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: positive_integral_def le_fun_def)
-lemma positive_integral_not_MInfty[simp]: "integral\<^isup>P M f \<noteq> -\<infinity>"
+lemma positive_integral_not_MInfty[simp]: "integral\<^sup>P M f \<noteq> -\<infinity>"
using positive_integral_positive[of M f] by auto
lemma positive_integral_def_finite:
- "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^isup>S M g)"
+ "integral\<^sup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^sup>S M g)"
(is "_ = SUPR ?A ?f")
unfolding positive_integral_def
proof (safe intro!: antisym SUP_least)
@@ -794,7 +794,7 @@
apply (safe intro!: simple_function_max simple_function_If)
apply (force simp: max_def le_fun_def split: split_if_asm)+
done
- show "integral\<^isup>S M g \<le> SUPR ?A ?f"
+ show "integral\<^sup>S M g \<le> SUPR ?A ?f"
proof cases
have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
assume "(emeasure M) ?G = 0"
@@ -805,7 +805,7 @@
(auto simp: max_def intro!: simple_function_If)
next
assume \<mu>_G: "(emeasure M) ?G \<noteq> 0"
- have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
+ have "SUPR ?A (integral\<^sup>S M) = \<infinity>"
proof (intro SUP_PInfty)
fix n :: nat
let ?y = "ereal (real n) / (if (emeasure M) ?G = \<infinity> then 1 else (emeasure M) ?G)"
@@ -813,12 +813,12 @@
then have "?g ?y \<in> ?A" by (rule g_in_A)
have "real n \<le> ?y * (emeasure M) ?G"
using \<mu>_G \<mu>_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps)
- also have "\<dots> = (\<integral>\<^isup>Sx. ?y * indicator ?G x \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>Sx. ?y * indicator ?G x \<partial>M)"
using `0 \<le> ?y` `?g ?y \<in> ?A` gM
by (subst simple_integral_cmult_indicator) auto
- also have "\<dots> \<le> integral\<^isup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
+ also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using `?g ?y \<in> ?A` gM
by (intro simple_integral_mono) auto
- finally show "\<exists>i\<in>?A. real n \<le> integral\<^isup>S M i"
+ finally show "\<exists>i\<in>?A. real n \<le> integral\<^sup>S M i"
using `?g ?y \<in> ?A` by blast
qed
then show ?thesis by simp
@@ -826,7 +826,7 @@
qed (auto intro: SUP_upper)
lemma positive_integral_mono_AE:
- assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
+ assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>P M u \<le> integral\<^sup>P M v"
unfolding positive_integral_def
proof (safe intro!: SUP_mono)
fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
@@ -844,34 +844,34 @@
by (auto simp: max_def split: split_if_asm)
qed simp }
then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI)
- moreover have "integral\<^isup>S M n \<le> integral\<^isup>S M ?n"
+ moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n"
using ae_N N n by (auto intro!: simple_integral_mono_AE)
- ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^isup>S M n \<le> integral\<^isup>S M m"
+ ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
by force
qed
lemma positive_integral_mono:
- "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
+ "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>P M u \<le> integral\<^sup>P M v"
by (auto intro: positive_integral_mono_AE)
lemma positive_integral_cong_AE:
- "AE x in M. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
+ "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>P M u = integral\<^sup>P M v"
by (auto simp: eq_iff intro!: positive_integral_mono_AE)
lemma positive_integral_cong:
- "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
+ "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>P M u = integral\<^sup>P M v"
by (auto intro: positive_integral_cong_AE)
lemma positive_integral_eq_simple_integral:
- assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+ assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^sup>P M f = integral\<^sup>S M f"
proof -
let ?f = "\<lambda>x. f x * indicator (space M) x"
have f': "simple_function M ?f" using f by auto
with f(2) have [simp]: "max 0 \<circ> ?f = ?f"
by (auto simp: fun_eq_iff max_def split: split_indicator)
- have "integral\<^isup>P M ?f \<le> integral\<^isup>S M ?f" using f'
+ have "integral\<^sup>P M ?f \<le> integral\<^sup>S M ?f" using f'
by (force intro!: SUP_least simple_integral_mono simp: le_fun_def positive_integral_def)
- moreover have "integral\<^isup>S M ?f \<le> integral\<^isup>P M ?f"
+ moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>P M ?f"
unfolding positive_integral_def
using f' by (auto intro!: SUP_upper)
ultimately show ?thesis
@@ -879,10 +879,10 @@
qed
lemma positive_integral_eq_simple_integral_AE:
- assumes f: "simple_function M f" "AE x in M. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+ assumes f: "simple_function M f" "AE x in M. 0 \<le> f x" shows "integral\<^sup>P M f = integral\<^sup>S M f"
proof -
have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max)
- with f have "integral\<^isup>P M f = integral\<^isup>S M (\<lambda>x. max 0 (f x))"
+ with f have "integral\<^sup>P M f = integral\<^sup>S M (\<lambda>x. max 0 (f x))"
by (simp cong: positive_integral_cong_AE simple_integral_cong_AE
add: positive_integral_eq_simple_integral)
with assms show ?thesis
@@ -892,11 +892,11 @@
lemma positive_integral_SUP_approx:
assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
- shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
+ shows "integral\<^sup>S M u \<le> (SUP i. integral\<^sup>P M (f i))" (is "_ \<le> ?S")
proof (rule ereal_le_mult_one_interval)
- have "0 \<le> (SUP i. integral\<^isup>P M (f i))"
+ have "0 \<le> (SUP i. integral\<^sup>P M (f i))"
using f(3) by (auto intro!: SUP_upper2 positive_integral_positive)
- then show "(SUP i. integral\<^isup>P M (f i)) \<noteq> -\<infinity>" by auto
+ then show "(SUP i. integral\<^sup>P M (f i)) \<noteq> -\<infinity>" by auto
have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>"
using u(3) by auto
fix a :: ereal assume "0 < a" "a < 1"
@@ -945,43 +945,43 @@
then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp
qed
- have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
+ have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))"
unfolding simple_integral_indicator[OF B `simple_function M u`]
proof (subst SUPR_ereal_setsum, safe)
fix x n assume "x \<in> space M"
with u_range show "incseq (\<lambda>i. u x * (emeasure M) (?B' (u x) i))" "\<And>i. 0 \<le> u x * (emeasure M) (?B' (u x) i)"
using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
next
- show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))"
+ show "integral\<^sup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))"
using measure_conv u_range B_u unfolding simple_integral_def
by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric])
qed
moreover
- have "a * (SUP i. integral\<^isup>S M (?uB i)) \<le> ?S"
+ have "a * (SUP i. integral\<^sup>S M (?uB i)) \<le> ?S"
apply (subst SUPR_ereal_cmult[symmetric])
proof (safe intro!: SUP_mono bexI)
fix i
- have "a * integral\<^isup>S M (?uB i) = (\<integral>\<^isup>Sx. a * ?uB i x \<partial>M)"
+ have "a * integral\<^sup>S M (?uB i) = (\<integral>\<^sup>Sx. a * ?uB i x \<partial>M)"
using B `simple_function M u` u_range
by (subst simple_integral_mult) (auto split: split_indicator)
- also have "\<dots> \<le> integral\<^isup>P M (f i)"
+ also have "\<dots> \<le> integral\<^sup>P M (f i)"
proof -
have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B `0 < a` u(1) by auto
show ?thesis using f(3) * u_range `0 < a`
by (subst positive_integral_eq_simple_integral[symmetric])
(auto intro!: positive_integral_mono split: split_indicator)
qed
- finally show "a * integral\<^isup>S M (?uB i) \<le> integral\<^isup>P M (f i)"
+ finally show "a * integral\<^sup>S M (?uB i) \<le> integral\<^sup>P M (f i)"
by auto
next
- fix i show "0 \<le> \<integral>\<^isup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
+ fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range
by (intro simple_integral_positive) (auto split: split_indicator)
qed (insert `0 < a`, auto)
- ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
+ ultimately show "a * integral\<^sup>S M u \<le> ?S" by simp
qed
lemma incseq_positive_integral:
- assumes "incseq f" shows "incseq (\<lambda>i. integral\<^isup>P M (f i))"
+ assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>P M (f i))"
proof -
have "\<And>i x. f i x \<le> f (Suc i) x"
using assms by (auto dest!: incseq_SucD simp: le_fun_def)
@@ -992,19 +992,19 @@
text {* Beppo-Levi monotone convergence theorem *}
lemma positive_integral_monotone_convergence_SUP:
assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
- shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+ shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>P M (f i))"
proof (rule antisym)
- show "(SUP j. integral\<^isup>P M (f j)) \<le> (\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M)"
+ show "(SUP j. integral\<^sup>P M (f j)) \<le> (\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M)"
by (auto intro!: SUP_least SUP_upper positive_integral_mono)
next
- show "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^isup>P M (f j))"
+ show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^sup>P M (f j))"
unfolding positive_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
proof (safe intro!: SUP_least)
fix g assume g: "simple_function M g"
and "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
moreover then have "\<And>x. 0 \<le> (SUP i. f i x)" and g': "g`space M \<subseteq> {0..<\<infinity>}"
using f by (auto intro!: SUP_upper2)
- ultimately show "integral\<^isup>S M g \<le> (SUP j. integral\<^isup>P M (f j))"
+ ultimately show "integral\<^sup>S M g \<le> (SUP j. integral\<^sup>P M (f j))"
by (intro positive_integral_SUP_approx[OF f g _ g'])
(auto simp: le_fun_def max_def)
qed
@@ -1012,16 +1012,16 @@
lemma positive_integral_monotone_convergence_SUP_AE:
assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+ shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>P M (f i))"
proof -
from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
by (simp add: AE_all_countable)
from this[THEN AE_E] guess N . note N = this
let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
- then have "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP i. ?f i x) \<partial>M)"
+ then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)"
by (auto intro!: positive_integral_cong_AE)
- also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. ?f i x \<partial>M))"
+ also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))"
proof (rule positive_integral_monotone_convergence_SUP)
show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
{ fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
@@ -1029,34 +1029,34 @@
fix x show "0 \<le> ?f i x"
using N(1) by auto }
qed
- also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. f i x \<partial>M))"
+ also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
using f_eq by (force intro!: arg_cong[where f="SUPR UNIV"] positive_integral_cong_AE ext)
finally show ?thesis .
qed
lemma positive_integral_monotone_convergence_SUP_AE_incseq:
assumes f: "incseq f" "\<And>i. AE x in M. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
- shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
+ shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>P M (f i))"
using f[unfolded incseq_Suc_iff le_fun_def]
by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel])
auto
lemma positive_integral_monotone_convergence_simple:
assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
- shows "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+ shows "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1)
f(3)[THEN borel_measurable_simple_function] f(2)]
by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPR UNIV"] ext)
lemma positive_integral_max_0:
- "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M) = integral\<^isup>P M f"
+ "(\<integral>\<^sup>+x. max 0 (f x) \<partial>M) = integral\<^sup>P M f"
by (simp add: le_fun_def positive_integral_def)
lemma positive_integral_cong_pos:
assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
- shows "integral\<^isup>P M f = integral\<^isup>P M g"
+ shows "integral\<^sup>P M f = integral\<^sup>P M g"
proof -
- have "integral\<^isup>P M (\<lambda>x. max 0 (f x)) = integral\<^isup>P M (\<lambda>x. max 0 (g x))"
+ have "integral\<^sup>P M (\<lambda>x. max 0 (f x)) = integral\<^sup>P M (\<lambda>x. max 0 (g x))"
proof (intro positive_integral_cong)
fix x assume "x \<in> space M"
from assms[OF this] show "max 0 (f x) = max 0 (g x)"
@@ -1069,12 +1069,12 @@
assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
- shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))"
+ shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))"
(is "SUPR _ ?F = SUPR _ ?G")
proof -
- have "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
+ have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
using f by (rule positive_integral_monotone_convergence_simple)
- also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. g i x) \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)"
unfolding eq[THEN positive_integral_cong_AE] ..
also have "\<dots> = (SUP i. ?G i)"
using g by (rule positive_integral_monotone_convergence_simple[symmetric])
@@ -1082,14 +1082,14 @@
qed
lemma positive_integral_const[simp]:
- "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * (emeasure M) (space M)"
+ "0 \<le> c \<Longrightarrow> (\<integral>\<^sup>+ x. c \<partial>M) = c * (emeasure M) (space M)"
by (subst positive_integral_eq_simple_integral) auto
lemma positive_integral_linear:
assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
- shows "(\<integral>\<^isup>+ x. a * f x + g x \<partial>M) = a * integral\<^isup>P M f + integral\<^isup>P M g"
- (is "integral\<^isup>P M ?L = _")
+ shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>P M f + integral\<^sup>P M g"
+ (is "integral\<^sup>P M ?L = _")
proof -
from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
@@ -1101,17 +1101,17 @@
from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this
- have inc: "incseq (\<lambda>i. a * integral\<^isup>S M (u i))" "incseq (\<lambda>i. integral\<^isup>S M (v i))"
+ have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
using u v `0 \<le> a`
by (auto simp: incseq_Suc_iff le_fun_def
intro!: add_mono ereal_mult_left_mono simple_integral_mono)
- have pos: "\<And>i. 0 \<le> integral\<^isup>S M (u i)" "\<And>i. 0 \<le> integral\<^isup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^isup>S M (u i)"
+ have pos: "\<And>i. 0 \<le> integral\<^sup>S M (u i)" "\<And>i. 0 \<le> integral\<^sup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^sup>S M (u i)"
using u v `0 \<le> a` by (auto simp: simple_integral_positive)
- { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \<noteq> -\<infinity>" "integral\<^isup>S M (v i) \<noteq> -\<infinity>"
+ { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \<noteq> -\<infinity>" "integral\<^sup>S M (v i) \<noteq> -\<infinity>"
by (auto split: split_if_asm) }
note not_MInf = this
- have l': "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))"
+ have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))"
proof (rule SUP_simple_integral_sequences[OF l(3,6,2)])
show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)"
using u v `0 \<le> a` unfolding incseq_Suc_iff le_fun_def
@@ -1128,9 +1128,9 @@
unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg)
qed
- also have "\<dots> = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))"
+ also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext)
- finally have "(\<integral>\<^isup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^isup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+x. max 0 (g x) \<partial>M)"
+ finally have "(\<integral>\<^sup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^sup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+x. max 0 (g x) \<partial>M)"
unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric]
unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric]
apply (subst SUPR_ereal_cmult[symmetric, OF pos(1) `0 \<le> a`])
@@ -1140,11 +1140,11 @@
lemma positive_integral_cmult:
assumes f: "f \<in> borel_measurable M" "0 \<le> c"
- shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
+ shows "(\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>P M f"
proof -
have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
by (auto split: split_max simp: ereal_zero_le_0_iff)
- have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. c * f x \<partial>M) = (\<integral>\<^sup>+ x. c * max 0 (f x) \<partial>M)"
by (simp add: positive_integral_max_0)
then show ?thesis
using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f
@@ -1153,24 +1153,24 @@
lemma positive_integral_multc:
assumes "f \<in> borel_measurable M" "0 \<le> c"
- shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
+ shows "(\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>P M f * c"
unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
lemma positive_integral_indicator[simp]:
- "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = (emeasure M) A"
+ "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
by (subst positive_integral_eq_simple_integral)
(auto simp: simple_integral_indicator)
lemma positive_integral_cmult_indicator:
- "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A"
+ "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A"
by (subst positive_integral_eq_simple_integral)
(auto simp: simple_function_indicator simple_integral_indicator)
lemma positive_integral_indicator':
assumes [measurable]: "A \<inter> space M \<in> sets M"
- shows "(\<integral>\<^isup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)"
+ shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)"
proof -
- have "(\<integral>\<^isup>+ x. indicator A x \<partial>M) = (\<integral>\<^isup>+ x. indicator (A \<inter> space M) x \<partial>M)"
+ have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)"
by (intro positive_integral_cong) (simp split: split_indicator)
also have "\<dots> = emeasure M (A \<inter> space M)"
by simp
@@ -1180,15 +1180,15 @@
lemma positive_integral_add:
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
- shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
+ shows "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>P M f + integral\<^sup>P M g"
proof -
have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg)
- have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = (\<integral>\<^sup>+ x. max 0 (f x + g x) \<partial>M)"
by (simp add: positive_integral_max_0)
- also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
unfolding ae[THEN positive_integral_cong_AE] ..
- also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+ x. max 0 (g x) \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (g x) \<partial>M)"
using positive_integral_linear[of "\<lambda>x. max 0 (f x)" _ 1 "\<lambda>x. max 0 (g x)"] f g
by auto
finally show ?thesis
@@ -1197,7 +1197,7 @@
lemma positive_integral_setsum:
assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x in M. 0 \<le> f i x"
- shows "(\<integral>\<^isup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>P M (f i))"
+ shows "(\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>P M (f i))"
proof cases
assume f: "finite P"
from assms have "AE x in M. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
@@ -1214,17 +1214,17 @@
lemma positive_integral_Markov_inequality:
assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c"
- shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+ shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
(is "(emeasure M) ?A \<le> _ * ?PI")
proof -
have "?A \<in> sets M"
using `A \<in> sets M` u by auto
- hence "(emeasure M) ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
+ hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
using positive_integral_indicator by simp
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
+ also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
by (auto intro!: positive_integral_mono_AE
simp: indicator_def ereal_zero_le_0_iff)
- also have "\<dots> = c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+ also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
using assms
by (auto intro!: positive_integral_cmult simp: ereal_zero_le_0_iff)
finally show ?thesis .
@@ -1232,7 +1232,7 @@
lemma positive_integral_noteq_infinite:
assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
- and "integral\<^isup>P M g \<noteq> \<infinity>"
+ and "integral\<^sup>P M g \<noteq> \<infinity>"
shows "AE x in M. g x \<noteq> \<infinity>"
proof (rule ccontr)
assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
@@ -1241,19 +1241,19 @@
moreover have "0 \<le> (emeasure M) {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
ultimately have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
- also have "\<dots> \<le> (\<integral>\<^isup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
using g by (subst positive_integral_cmult_indicator) auto
- also have "\<dots> \<le> integral\<^isup>P M g"
+ also have "\<dots> \<le> integral\<^sup>P M g"
using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def)
- finally show False using `integral\<^isup>P M g \<noteq> \<infinity>` by auto
+ finally show False using `integral\<^sup>P M g \<noteq> \<infinity>` by auto
qed
lemma positive_integral_diff:
assumes f: "f \<in> borel_measurable M"
and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
- and fin: "integral\<^isup>P M g \<noteq> \<infinity>"
+ and fin: "integral\<^sup>P M g \<noteq> \<infinity>"
and mono: "AE x in M. g x \<le> f x"
- shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
+ shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>P M f - integral\<^sup>P M g"
proof -
have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x in M. 0 \<le> f x - g x"
using assms by (auto intro: ereal_diff_positive)
@@ -1263,28 +1263,28 @@
note * = this
then have "AE x in M. f x = f x - g x + g x"
using mono positive_integral_noteq_infinite[OF g fin] assms by auto
- then have **: "integral\<^isup>P M f = (\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g"
+ then have **: "integral\<^sup>P M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>P M g"
unfolding positive_integral_add[OF diff g, symmetric]
by (rule positive_integral_cong_AE)
show ?thesis unfolding **
using fin positive_integral_positive[of M g]
- by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
+ by (cases rule: ereal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>P M g"]) auto
qed
lemma positive_integral_suminf:
assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> f i x"
- shows "(\<integral>\<^isup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^isup>P M (f i))"
+ shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>P M (f i))"
proof -
have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
using assms by (auto simp: AE_all_countable)
- have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
+ have "(\<Sum>i. integral\<^sup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>P M (f i))"
using positive_integral_positive by (rule suminf_ereal_eq_SUPR)
- also have "\<dots> = (SUP n. \<integral>\<^isup>+x. (\<Sum>i<n. f i x) \<partial>M)"
+ also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
unfolding positive_integral_setsum[OF f] ..
- also have "\<dots> = \<integral>\<^isup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
+ also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
(elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
- also have "\<dots> = \<integral>\<^isup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
+ also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUPR)
finally show ?thesis by simp
qed
@@ -1293,24 +1293,24 @@
lemma positive_integral_lim_INF:
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> u i x"
- shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
+ shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>P M (u n))"
proof -
have pos: "AE x in M. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
- have "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
- (SUP n. \<integral>\<^isup>+ x. (INF i:{n..}. u i x) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
+ (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
unfolding liminf_SUPR_INFI using pos u
by (intro positive_integral_monotone_convergence_SUP_AE)
(elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
- also have "\<dots> \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
+ also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>P M (u n))"
unfolding liminf_SUPR_INFI
by (auto intro!: SUP_mono exI INF_greatest positive_integral_mono INF_lower)
finally show ?thesis .
qed
lemma positive_integral_null_set:
- assumes "N \<in> null_sets M" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
+ assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0"
proof -
- have "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
+ have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
proof (intro positive_integral_cong_AE AE_I)
show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
by (auto simp: indicator_def)
@@ -1322,22 +1322,22 @@
lemma positive_integral_0_iff:
assumes u: "u \<in> borel_measurable M" and pos: "AE x in M. 0 \<le> u x"
- shows "integral\<^isup>P M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
+ shows "integral\<^sup>P M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
(is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
proof -
- have u_eq: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
+ have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>P M u"
by (auto intro!: positive_integral_cong simp: indicator_def)
show ?thesis
proof
assume "(emeasure M) ?A = 0"
with positive_integral_null_set[of ?A M u] u
- show "integral\<^isup>P M u = 0" by (simp add: u_eq null_sets_def)
+ show "integral\<^sup>P M u = 0" by (simp add: u_eq null_sets_def)
next
{ fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def)
then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) }
note gt_1 = this
- assume *: "integral\<^isup>P M u = 0"
+ assume *: "integral\<^sup>P M u = 0"
let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
proof -
@@ -1387,12 +1387,12 @@
lemma positive_integral_0_iff_AE:
assumes u: "u \<in> borel_measurable M"
- shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)"
+ shows "integral\<^sup>P M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)"
proof -
have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
using u by auto
from positive_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
- have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)"
+ have "integral\<^sup>P M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)"
unfolding positive_integral_max_0
using AE_iff_null[OF sets] u by auto
also have "\<dots> \<longleftrightarrow> (AE x in M. u x \<le> 0)" by (auto split: split_max)
@@ -1400,18 +1400,18 @@
qed
lemma AE_iff_positive_integral:
- "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^isup>P M (indicator {x. \<not> P x}) = 0"
+ "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>P M (indicator {x. \<not> P x}) = 0"
by (subst positive_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def
sets.sets_Collect_neg indicator_def[abs_def] measurable_If)
lemma positive_integral_const_If:
- "(\<integral>\<^isup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)"
+ "(\<integral>\<^sup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)"
by (auto intro!: positive_integral_0_iff_AE[THEN iffD2])
lemma positive_integral_subalgebra:
assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
- shows "integral\<^isup>P N f = integral\<^isup>P M f"
+ shows "integral\<^sup>P N f = integral\<^sup>P M f"
proof -
have [simp]: "\<And>f :: 'a \<Rightarrow> ereal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
using N by (auto simp: measurable_def)
@@ -1429,7 +1429,7 @@
lemma positive_integral_nat_function:
fixes f :: "'a \<Rightarrow> nat"
assumes "f \<in> measurable M (count_space UNIV)"
- shows "(\<integral>\<^isup>+x. ereal (of_nat (f x)) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
+ shows "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
proof -
def F \<equiv> "\<lambda>i. {x\<in>space M. i < f x}"
with assms have [measurable]: "\<And>i. F i \<in> sets M"
@@ -1444,7 +1444,7 @@
using `x \<in> space M` by (simp add: one_ereal_def F_def)
ultimately have "ereal(of_nat(f x)) = (\<Sum>i. indicator (F i) x)"
by (simp add: sums_iff) }
- then have "(\<integral>\<^isup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
+ then have "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
by (simp cong: positive_integral_cong)
also have "\<dots> = (\<Sum>i. emeasure M (F i))"
by (simp add: positive_integral_suminf)
@@ -1456,7 +1456,7 @@
definition integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
"integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
- (\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
+ (\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
lemma borel_measurable_integrable[measurable_dest]:
"integrable M f \<Longrightarrow> f \<in> borel_measurable M"
@@ -1464,11 +1464,11 @@
lemma integrableD[dest]:
assumes "integrable M f"
- shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
+ shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
using assms unfolding integrable_def by auto
-definition lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> real" ("integral\<^isup>L") where
- "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. ereal (- f x) \<partial>M))"
+definition lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> real" ("integral\<^sup>L") where
+ "integral\<^sup>L M f = real ((\<integral>\<^sup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^sup>+ x. ereal (- f x) \<partial>M))"
syntax
"_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
@@ -1479,22 +1479,22 @@
lemma integrableE:
assumes "integrable M f"
obtains r q where
- "(\<integral>\<^isup>+x. ereal (f x)\<partial>M) = ereal r"
- "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M) = ereal q"
- "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
+ "(\<integral>\<^sup>+x. ereal (f x)\<partial>M) = ereal r"
+ "(\<integral>\<^sup>+x. ereal (-f x)\<partial>M) = ereal q"
+ "f \<in> borel_measurable M" "integral\<^sup>L M f = r - q"
using assms unfolding integrable_def lebesgue_integral_def
using positive_integral_positive[of M "\<lambda>x. ereal (f x)"]
using positive_integral_positive[of M "\<lambda>x. ereal (-f x)"]
- by (cases rule: ereal2_cases[of "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. ereal (f x)\<partial>M)"]) auto
+ by (cases rule: ereal2_cases[of "(\<integral>\<^sup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ereal (f x)\<partial>M)"]) auto
lemma integral_cong:
assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "integral\<^isup>L M f = integral\<^isup>L M g"
+ shows "integral\<^sup>L M f = integral\<^sup>L M g"
using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def)
lemma integral_cong_AE:
assumes cong: "AE x in M. f x = g x"
- shows "integral\<^isup>L M f = integral\<^isup>L M g"
+ shows "integral\<^sup>L M f = integral\<^sup>L M g"
proof -
have *: "AE x in M. ereal (f x) = ereal (g x)"
"AE x in M. ereal (- f x) = ereal (- g x)" using cong by auto
@@ -1507,8 +1507,8 @@
assumes "AE x in M. f x = g x"
shows "integrable M f = integrable M g"
proof -
- have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (g x) \<partial>M)"
- "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (- g x) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ereal (g x) \<partial>M)"
+ "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) = (\<integral>\<^sup>+ x. ereal (- g x) \<partial>M)"
using assms by (auto intro!: positive_integral_cong_AE)
with assms show ?thesis
by (auto simp: integrable_def)
@@ -1521,7 +1521,7 @@
lemma integral_mono_AE:
assumes fg: "integrable M f" "integrable M g"
and mono: "AE t in M. f t \<le> g t"
- shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
+ shows "integral\<^sup>L M f \<le> integral\<^sup>L M g"
proof -
have "AE x in M. ereal (f x) \<le> ereal (g x)"
using mono by auto
@@ -1534,35 +1534,35 @@
lemma integral_mono:
assumes "integrable M f" "integrable M g" "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
- shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
+ shows "integral\<^sup>L M f \<le> integral\<^sup>L M g"
using assms by (auto intro: integral_mono_AE)
lemma positive_integral_eq_integral:
assumes f: "integrable M f"
assumes nonneg: "AE x in M. 0 \<le> f x"
- shows "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = integral\<^isup>L M f"
+ shows "(\<integral>\<^sup>+ x. ereal (f x) \<partial>M) = integral\<^sup>L M f"
proof -
- have "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
+ have "(\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
using nonneg by (intro positive_integral_cong_AE) (auto split: split_max)
with f positive_integral_positive show ?thesis
- by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>M")
+ by (cases "\<integral>\<^sup>+ x. ereal (f x) \<partial>M")
(auto simp add: lebesgue_integral_def positive_integral_max_0 integrable_def)
qed
lemma integral_eq_positive_integral:
assumes f: "\<And>x. 0 \<le> f x"
- shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
+ shows "integral\<^sup>L M f = real (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"
proof -
{ fix x have "max 0 (ereal (- f x)) = 0" using f[of x] by (simp split: split_max) }
- then have "0 = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)" by simp
- also have "\<dots> = (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
+ then have "0 = (\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>M)" by simp
+ also have "\<dots> = (\<integral>\<^sup>+ x. ereal (- f x) \<partial>M)" unfolding positive_integral_max_0 ..
finally show ?thesis
unfolding lebesgue_integral_def by simp
qed
lemma integral_minus[intro, simp]:
assumes "integrable M f"
- shows "integrable M (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
+ shows "integrable M (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^sup>L M f"
using assms by (auto simp: integrable_def lebesgue_integral_def)
lemma integral_minus_iff[simp]:
@@ -1577,7 +1577,7 @@
lemma integral_of_positive_diff:
assumes integrable: "integrable M u" "integrable M v"
and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
- shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
+ shows "integrable M f" and "integral\<^sup>L M f = integral\<^sup>L M u - integral\<^sup>L M v"
proof -
let ?f = "\<lambda>x. max 0 (ereal (f x))"
let ?mf = "\<lambda>x. max 0 (ereal (- f x))"
@@ -1592,9 +1592,9 @@
"f \<in> borel_measurable M"
by (auto simp: f_def[symmetric] integrable_def)
- have "(\<integral>\<^isup>+ x. ereal (u x - v x) \<partial>M) \<le> integral\<^isup>P M ?u"
+ have "(\<integral>\<^sup>+ x. ereal (u x - v x) \<partial>M) \<le> integral\<^sup>P M ?u"
using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
- moreover have "(\<integral>\<^isup>+ x. ereal (v x - u x) \<partial>M) \<le> integral\<^isup>P M ?v"
+ moreover have "(\<integral>\<^sup>+ x. ereal (v x - u x) \<partial>M) \<le> integral\<^sup>P M ?v"
using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0)
ultimately show f: "integrable M f"
using `integrable M u` `integrable M v` `f \<in> borel_measurable M`
@@ -1602,13 +1602,13 @@
have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
unfolding f_def using pos by (simp split: split_max)
- then have "(\<integral>\<^isup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^isup>+ x. ?v x + ?f x \<partial>M)" by simp
- then have "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) =
- real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)"
+ then have "(\<integral>\<^sup>+ x. ?u x + ?mf x \<partial>M) = (\<integral>\<^sup>+ x. ?v x + ?f x \<partial>M)" by simp
+ then have "real (integral\<^sup>P M ?u + integral\<^sup>P M ?mf) =
+ real (integral\<^sup>P M ?v + integral\<^sup>P M ?f)"
using positive_integral_add[OF u_borel _ mf_borel]
using positive_integral_add[OF v_borel _ f_borel]
by auto
- then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
+ then show "integral\<^sup>L M f = integral\<^sup>L M u - integral\<^sup>L M v"
unfolding positive_integral_max_0
unfolding pos[THEN integral_eq_positive_integral]
using integrable f by (auto elim!: integrableE)
@@ -1617,7 +1617,7 @@
lemma integral_linear:
assumes "integrable M f" "integrable M g" and "0 \<le> a"
shows "integrable M (\<lambda>t. a * f t + g t)"
- and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
+ and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^sup>L M f + integral\<^sup>L M g" (is ?EQ)
proof -
let ?f = "\<lambda>x. max 0 (ereal (f x))"
let ?g = "\<lambda>x. max 0 (ereal (g x))"
@@ -1627,11 +1627,11 @@
let ?n = "\<lambda>t. max 0 (- (a * f t)) + max 0 (- g t)"
from assms have linear:
- "(\<integral>\<^isup>+ x. ereal a * ?f x + ?g x \<partial>M) = ereal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g"
- "(\<integral>\<^isup>+ x. ereal a * ?mf x + ?mg x \<partial>M) = ereal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg"
+ "(\<integral>\<^sup>+ x. ereal a * ?f x + ?g x \<partial>M) = ereal a * integral\<^sup>P M ?f + integral\<^sup>P M ?g"
+ "(\<integral>\<^sup>+ x. ereal a * ?mf x + ?mg x \<partial>M) = ereal a * integral\<^sup>P M ?mf + integral\<^sup>P M ?mg"
by (auto intro!: positive_integral_linear simp: integrable_def)
- have *: "(\<integral>\<^isup>+x. ereal (- ?p x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- ?n x) \<partial>M) = 0"
+ have *: "(\<integral>\<^sup>+x. ereal (- ?p x) \<partial>M) = 0" "(\<integral>\<^sup>+x. ereal (- ?n x) \<partial>M) = 0"
using `0 \<le> a` assms by (auto simp: positive_integral_0_iff_AE integrable_def)
have **: "\<And>x. ereal a * ?f x + ?g x = max 0 (ereal (?p x))"
"\<And>x. ereal a * ?mf x + ?mg x = max 0 (ereal (?n x))"
@@ -1653,7 +1653,7 @@
lemma integral_add[simp, intro]:
assumes "integrable M f" "integrable M g"
shows "integrable M (\<lambda>t. f t + g t)"
- and "(\<integral> t. f t + g t \<partial>M) = integral\<^isup>L M f + integral\<^isup>L M g"
+ and "(\<integral> t. f t + g t \<partial>M) = integral\<^sup>L M f + integral\<^sup>L M g"
using assms integral_linear[where a=1] by auto
lemma integral_zero[simp, intro]:
@@ -1662,27 +1662,27 @@
by auto
lemma lebesgue_integral_uminus:
- "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
+ "(\<integral>x. - f x \<partial>M) = - integral\<^sup>L M f"
unfolding lebesgue_integral_def by simp
lemma lebesgue_integral_cmult_nonneg:
assumes f: "f \<in> borel_measurable M" and "0 \<le> c"
- shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^isup>L M f"
+ shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^sup>L M f"
proof -
- { have "real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x)))) =
- real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (f x))))"
+ { have "real (ereal c * integral\<^sup>P M (\<lambda>x. max 0 (ereal (f x)))) =
+ real (integral\<^sup>P M (\<lambda>x. ereal c * max 0 (ereal (f x))))"
using f `0 \<le> c` by (subst positive_integral_cmult) auto
- also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (c * f x))))"
+ also have "\<dots> = real (integral\<^sup>P M (\<lambda>x. max 0 (ereal (c * f x))))"
using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def zero_le_mult_iff)
- finally have "real (integral\<^isup>P M (\<lambda>x. ereal (c * f x))) = c * real (integral\<^isup>P M (\<lambda>x. ereal (f x)))"
+ finally have "real (integral\<^sup>P M (\<lambda>x. ereal (c * f x))) = c * real (integral\<^sup>P M (\<lambda>x. ereal (f x)))"
by (simp add: positive_integral_max_0) }
moreover
- { have "real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x)))) =
- real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (- f x))))"
+ { have "real (ereal c * integral\<^sup>P M (\<lambda>x. max 0 (ereal (- f x)))) =
+ real (integral\<^sup>P M (\<lambda>x. ereal c * max 0 (ereal (- f x))))"
using f `0 \<le> c` by (subst positive_integral_cmult) auto
- also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (- c * f x))))"
+ also have "\<dots> = real (integral\<^sup>P M (\<lambda>x. max 0 (ereal (- c * f x))))"
using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def mult_le_0_iff)
- finally have "real (integral\<^isup>P M (\<lambda>x. ereal (- c * f x))) = c * real (integral\<^isup>P M (\<lambda>x. ereal (- f x)))"
+ finally have "real (integral\<^sup>P M (\<lambda>x. ereal (- c * f x))) = c * real (integral\<^sup>P M (\<lambda>x. ereal (- f x)))"
by (simp add: positive_integral_max_0) }
ultimately show ?thesis
by (simp add: lebesgue_integral_def field_simps)
@@ -1690,7 +1690,7 @@
lemma lebesgue_integral_cmult:
assumes f: "f \<in> borel_measurable M"
- shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^isup>L M f"
+ shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^sup>L M f"
proof (cases rule: linorder_le_cases)
assume "0 \<le> c" with f show ?thesis by (rule lebesgue_integral_cmult_nonneg)
next
@@ -1701,19 +1701,19 @@
qed
lemma lebesgue_integral_multc:
- "f \<in> borel_measurable M \<Longrightarrow> (\<integral>x. f x * c \<partial>M) = integral\<^isup>L M f * c"
+ "f \<in> borel_measurable M \<Longrightarrow> (\<integral>x. f x * c \<partial>M) = integral\<^sup>L M f * c"
using lebesgue_integral_cmult[of f M c] by (simp add: ac_simps)
lemma integral_multc:
- "integrable M f \<Longrightarrow> (\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
+ "integrable M f \<Longrightarrow> (\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
by (simp add: lebesgue_integral_multc)
lemma integral_cmult[simp, intro]:
assumes "integrable M f"
shows "integrable M (\<lambda>t. a * f t)" (is ?P)
- and "(\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f" (is ?I)
+ and "(\<integral> t. a * f t \<partial>M) = a * integral\<^sup>L M f" (is ?I)
proof -
- have "integrable M (\<lambda>t. a * f t) \<and> (\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f"
+ have "integrable M (\<lambda>t. a * f t) \<and> (\<integral> t. a * f t \<partial>M) = a * integral\<^sup>L M f"
proof (cases rule: le_cases)
assume "0 \<le> a" show ?thesis
using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
@@ -1731,19 +1731,19 @@
lemma integral_diff[simp, intro]:
assumes f: "integrable M f" and g: "integrable M g"
shows "integrable M (\<lambda>t. f t - g t)"
- and "(\<integral> t. f t - g t \<partial>M) = integral\<^isup>L M f - integral\<^isup>L M g"
+ and "(\<integral> t. f t - g t \<partial>M) = integral\<^sup>L M f - integral\<^sup>L M g"
using integral_add[OF f integral_minus(1)[OF g]]
unfolding diff_minus integral_minus(2)[OF g]
by auto
lemma integral_indicator[simp, intro]:
assumes "A \<in> sets M" and "(emeasure M) A \<noteq> \<infinity>"
- shows "integral\<^isup>L M (indicator A) = real (emeasure M A)" (is ?int)
+ shows "integral\<^sup>L M (indicator A) = real (emeasure M A)" (is ?int)
and "integrable M (indicator A)" (is ?able)
proof -
from `A \<in> sets M` have *:
"\<And>x. ereal (indicator A x) = indicator A x"
- "(\<integral>\<^isup>+x. ereal (- indicator A x) \<partial>M) = 0"
+ "(\<integral>\<^sup>+x. ereal (- indicator A x) \<partial>M) = 0"
by (auto split: split_indicator simp: positive_integral_0_iff_AE one_ereal_def)
show ?int ?able
using assms unfolding lebesgue_integral_def integrable_def
@@ -1768,7 +1768,7 @@
lemma integral_setsum[simp, intro]:
assumes "\<And>n. n \<in> S \<Longrightarrow> integrable M (f n)"
- shows "(\<integral>x. (\<Sum> i \<in> S. f i x) \<partial>M) = (\<Sum> i \<in> S. integral\<^isup>L M (f i))" (is "?int S")
+ shows "(\<integral>x. (\<Sum> i \<in> S. f i x) \<partial>M) = (\<Sum> i \<in> S. integral\<^sup>L M (f i))" (is "?int S")
and "integrable M (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
proof -
have "?int S \<and> ?I S"
@@ -1784,21 +1784,21 @@
assumes borel: "g \<in> borel_measurable M"
shows "integrable M g"
proof -
- have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
+ have "(\<integral>\<^sup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
+ also have "\<dots> \<le> (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono_AE)
also have "\<dots> < \<infinity>"
using `integrable M f` unfolding integrable_def by auto
- finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
+ finally have pos: "(\<integral>\<^sup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
- have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
by (auto intro!: positive_integral_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
+ also have "\<dots> \<le> (\<integral>\<^sup>+ x. ereal (f x) \<partial>M)"
using f by (auto intro!: positive_integral_mono_AE)
also have "\<dots> < \<infinity>"
using `integrable M f` unfolding integrable_def by auto
- finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
+ finally have neg: "(\<integral>\<^sup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
from neg pos borel show ?thesis
unfolding integrable_def by auto
@@ -1808,7 +1808,7 @@
assumes f[measurable]: "integrable M f"
shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
proof -
- from assms have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
+ from assms have *: "(\<integral>\<^sup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
"\<And>x. ereal \<bar>f x\<bar> = max 0 (ereal (f x)) + max 0 (ereal (- f x))"
by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max)
with assms show ?thesis
@@ -1819,10 +1819,10 @@
assumes borel: "f \<in> borel_measurable N"
and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
- and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
+ and "integral\<^sup>L N f = integral\<^sup>L M f" (is ?I)
proof -
- have "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M)"
- "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)"
+ have "(\<integral>\<^sup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^sup>+ x. max 0 (ereal (f x)) \<partial>M)"
+ "(\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^sup>+ x. max 0 (ereal (- f x)) \<partial>M)"
using borel by (auto intro!: positive_integral_subalgebra N)
moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
using assms unfolding measurable_def by auto
@@ -1831,9 +1831,9 @@
qed
lemma lebesgue_integral_nonneg:
- assumes ae: "(AE x in M. 0 \<le> f x)" shows "0 \<le> integral\<^isup>L M f"
+ assumes ae: "(AE x in M. 0 \<le> f x)" shows "0 \<le> integral\<^sup>L M f"
proof -
- have "(\<integral>\<^isup>+x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^isup>+x. 0 \<partial>M)"
+ have "(\<integral>\<^sup>+x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^sup>+x. 0 \<partial>M)"
using ae by (intro positive_integral_cong_AE) (auto simp: max_def)
then show ?thesis
by (auto simp: lebesgue_integral_def positive_integral_max_0
@@ -1866,9 +1866,9 @@
lemma integral_triangle_inequality:
assumes "integrable M f"
- shows "\<bar>integral\<^isup>L M f\<bar> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
+ shows "\<bar>integral\<^sup>L M f\<bar> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
proof -
- have "\<bar>integral\<^isup>L M f\<bar> = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto
+ have "\<bar>integral\<^sup>L M f\<bar> = max (integral\<^sup>L M f) (- integral\<^sup>L M f)" by auto
also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
using assms integral_minus(2)[of M f, symmetric]
by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
@@ -1876,21 +1876,21 @@
qed
lemma integrable_nonneg:
- assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^isup>+ x. f x \<partial>M) \<noteq> \<infinity>"
+ assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+ x. f x \<partial>M) \<noteq> \<infinity>"
shows "integrable M f"
unfolding integrable_def
proof (intro conjI f)
- have "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) = 0"
+ have "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) = 0"
using f by (subst positive_integral_0_iff_AE) auto
- then show "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>" by simp
+ then show "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>" by simp
qed
lemma integral_positive:
assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
- shows "0 \<le> integral\<^isup>L M f"
+ shows "0 \<le> integral\<^sup>L M f"
proof -
have "0 = (\<integral>x. 0 \<partial>M)" by auto
- also have "\<dots> \<le> integral\<^isup>L M f"
+ also have "\<dots> \<le> integral\<^sup>L M f"
using assms by (rule integral_mono[OF integral_zero(1)])
finally show ?thesis .
qed
@@ -1899,12 +1899,12 @@
assumes i: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
and pos: "\<And>i. AE x in M. 0 \<le> f i x"
and lim: "AE x in M. (\<lambda>i. f i x) ----> u x"
- and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+ and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) ----> x"
and u: "u \<in> borel_measurable M"
shows "integrable M u"
- and "integral\<^isup>L M u = x"
+ and "integral\<^sup>L M u = x"
proof -
- have "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M))"
+ have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ereal (f n x) \<partial>M))"
proof (subst positive_integral_monotone_convergence_SUP_AE[symmetric])
fix i
from mono pos show "AE x in M. ereal (f i x) \<le> ereal (f (Suc i) x) \<and> 0 \<le> ereal (f i x)"
@@ -1912,7 +1912,7 @@
show "(\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
using i by auto
next
- show "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = \<integral>\<^isup>+ x. (SUP i. ereal (f i x)) \<partial>M"
+ show "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ereal (f i x)) \<partial>M"
apply (rule positive_integral_cong_AE)
using lim mono
by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2])
@@ -1920,8 +1920,8 @@
also have "\<dots> = ereal x"
using mono i unfolding positive_integral_eq_integral[OF i pos]
by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim)
- finally have "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = ereal x" .
- moreover have "(\<integral>\<^isup>+ x. ereal (- u x) \<partial>M) = 0"
+ finally have "(\<integral>\<^sup>+ x. ereal (u x) \<partial>M) = ereal x" .
+ moreover have "(\<integral>\<^sup>+ x. ereal (- u x) \<partial>M) = 0"
proof (subst positive_integral_0_iff_AE)
show "(\<lambda>x. ereal (- u x)) \<in> borel_measurable M"
using u by auto
@@ -1932,17 +1932,17 @@
using incseq_le[of "\<lambda>n. f n x" "u x" 0] by (simp add: mono_def incseq_def)
qed
qed
- ultimately show "integrable M u" "integral\<^isup>L M u = x"
+ ultimately show "integrable M u" "integral\<^sup>L M u = x"
by (auto simp: integrable_def lebesgue_integral_def u)
qed
lemma integral_monotone_convergence:
assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
and lim: "AE x in M. (\<lambda>i. f i x) ----> u x"
- and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+ and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) ----> x"
and u: "u \<in> borel_measurable M"
shows "integrable M u"
- and "integral\<^isup>L M u = x"
+ and "integral\<^sup>L M u = x"
proof -
have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
using f by auto
@@ -1952,14 +1952,14 @@
using mono by (auto simp: field_simps mono_def le_fun_def)
have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x"
using lim by (auto intro!: tendsto_diff)
- have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^isup>L M (f 0)"
+ have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) ----> x - integral\<^sup>L M (f 0)"
using f ilim by (auto intro!: tendsto_diff)
have 6: "(\<lambda>x. u x - f 0 x) \<in> borel_measurable M"
using f[of 0] u by auto
note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5 6]
have "integrable M (\<lambda>x. (u x - f 0 x) + f 0 x)"
using diff(1) f by (rule integral_add(1))
- with diff(2) f show "integrable M u" "integral\<^isup>L M u = x"
+ with diff(2) f show "integrable M u" "integral\<^sup>L M u = x"
by auto
qed
@@ -1967,11 +1967,11 @@
assumes "integrable M f"
shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> (emeasure M) {x\<in>space M. f x \<noteq> 0} = 0"
proof -
- have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
+ have *: "(\<integral>\<^sup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
have "integrable M (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
hence "(\<lambda>x. ereal (\<bar>f x\<bar>)) \<in> borel_measurable M"
- "(\<integral>\<^isup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
+ "(\<integral>\<^sup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
from positive_integral_0_iff[OF this(1)] this(2)
show ?thesis unfolding lebesgue_integral_def *
using positive_integral_positive[of M "\<lambda>x. ereal \<bar>f x\<bar>"]
@@ -1980,14 +1980,14 @@
lemma positive_integral_PInf:
assumes f: "f \<in> borel_measurable M"
- and not_Inf: "integral\<^isup>P M f \<noteq> \<infinity>"
+ and not_Inf: "integral\<^sup>P M f \<noteq> \<infinity>"
shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
proof -
- have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+ have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets)
- also have "\<dots> \<le> integral\<^isup>P M (\<lambda>x. max 0 (f x))"
+ also have "\<dots> \<le> integral\<^sup>P M (\<lambda>x. max 0 (f x))"
by (auto intro!: positive_integral_mono simp: indicator_def max_def)
- finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
+ finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>P M f"
by (simp add: positive_integral_max_0)
moreover have "0 \<le> (emeasure M) (f -` {\<infinity>} \<inter> space M)"
by (rule emeasure_nonneg)
@@ -1996,7 +1996,7 @@
qed
lemma positive_integral_PInf_AE:
- assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
+ assumes "f \<in> borel_measurable M" "integral\<^sup>P M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
proof (rule AE_I)
show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
by (rule positive_integral_PInf[OF assms])
@@ -2006,16 +2006,16 @@
lemma simple_integral_PInf:
assumes "simple_function M f" "\<And>x. 0 \<le> f x"
- and "integral\<^isup>S M f \<noteq> \<infinity>"
+ and "integral\<^sup>S M f \<noteq> \<infinity>"
shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
proof (rule positive_integral_PInf)
show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
- show "integral\<^isup>P M f \<noteq> \<infinity>"
+ show "integral\<^sup>P M f \<noteq> \<infinity>"
using assms by (simp add: positive_integral_eq_simple_integral)
qed
lemma integral_real:
- "AE x in M. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
+ "AE x in M. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^sup>P M f) - real (integral\<^sup>P M (\<lambda>x. - f x))"
using assms unfolding lebesgue_integral_def
by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real)
@@ -2024,10 +2024,10 @@
and "(\<integral>x. a \<partial>M) = a * measure M (space M)"
proof -
{ fix a :: real assume "0 \<le> a"
- then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * (emeasure M) (space M)"
+ then have "(\<integral>\<^sup>+ x. ereal a \<partial>M) = ereal a * (emeasure M) (space M)"
by (subst positive_integral_const) auto
moreover
- from `0 \<le> a` have "(\<integral>\<^isup>+ x. ereal (-a) \<partial>M) = 0"
+ from `0 \<le> a` have "(\<integral>\<^sup>+ x. ereal (-a) \<partial>M) = 0"
by (subst positive_integral_0_iff_AE) auto
ultimately have "integrable M (\<lambda>x. a)" by (auto simp: integrable_def) }
note * = this
@@ -2055,22 +2055,22 @@
assumes int: "integrable M X" "integrable M Y"
assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x"
assumes gt: "AE x in M. X x \<le> Y x"
- shows "integral\<^isup>L M X < integral\<^isup>L M Y"
+ shows "integral\<^sup>L M X < integral\<^sup>L M Y"
proof -
- have "integral\<^isup>L M X \<le> integral\<^isup>L M Y"
+ have "integral\<^sup>L M X \<le> integral\<^sup>L M Y"
using gt int by (intro integral_mono_AE) auto
moreover
- have "integral\<^isup>L M X \<noteq> integral\<^isup>L M Y"
+ have "integral\<^sup>L M X \<noteq> integral\<^sup>L M Y"
proof
- assume eq: "integral\<^isup>L M X = integral\<^isup>L M Y"
- have "integral\<^isup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^isup>L M (\<lambda>x. Y x - X x)"
+ assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y"
+ have "integral\<^sup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^sup>L M (\<lambda>x. Y x - X x)"
using gt by (intro integral_cong_AE) auto
also have "\<dots> = 0"
using eq int by simp
finally have "(emeasure M) {x \<in> space M. Y x - X x \<noteq> 0} = 0"
using int by (simp add: integral_0_iff)
moreover
- have "(\<integral>\<^isup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^isup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
+ have "(\<integral>\<^sup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^sup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
using A by (intro positive_integral_mono_AE) auto
then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}"
using int A by (simp add: integrable_def)
@@ -2084,7 +2084,7 @@
lemma (in finite_measure) integral_less_AE_space:
assumes int: "integrable M X" "integrable M Y"
assumes gt: "AE x in M. X x < Y x" "(emeasure M) (space M) \<noteq> 0"
- shows "integral\<^isup>L M X < integral\<^isup>L M Y"
+ shows "integral\<^sup>L M X < integral\<^sup>L M Y"
using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
lemma integral_dominated_convergence:
@@ -2094,7 +2094,7 @@
and [measurable]: "u' \<in> borel_measurable M"
shows "integrable M u'"
and "(\<lambda>i. (\<integral>x. \<bar>u i x - u' x\<bar> \<partial>M)) ----> 0" (is "?lim_diff")
- and "(\<lambda>i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim)
+ and "(\<lambda>i. integral\<^sup>L M (u i)) ----> integral\<^sup>L M u'" (is ?lim)
proof -
have all_bound: "AE x in M. \<forall>j. \<bar>u j x\<bar> \<le> w x"
using bound by (auto simp: AE_all_countable)
@@ -2121,35 +2121,35 @@
finally show "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp
qed
- have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. ereal (?diff n x) \<partial>M) =
- (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
+ have PI_diff: "\<And>n. (\<integral>\<^sup>+ x. ereal (?diff n x) \<partial>M) =
+ (\<integral>\<^sup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
using diff w diff_less_2w w_pos
by (subst positive_integral_diff[symmetric])
(auto simp: integrable_def intro!: positive_integral_cong_AE)
have "integrable M (\<lambda>x. 2 * w x)"
using w by auto
- hence I2w_fin: "(\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
+ hence I2w_fin: "(\<integral>\<^sup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
borel_2w: "(\<lambda>x. ereal (2 * w x)) \<in> borel_measurable M"
unfolding integrable_def by auto
- have "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
+ have "limsup (\<lambda>n. \<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = 0" (is "limsup ?f = 0")
proof cases
- assume eq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
+ assume eq_0: "(\<integral>\<^sup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
{ fix n
- have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
+ have "?f n \<le> ?wx" (is "integral\<^sup>P M ?f' \<le> _")
using diff_less_2w unfolding positive_integral_max_0
by (intro positive_integral_mono_AE) auto
then have "?f n = 0"
using positive_integral_positive[of M ?f'] eq_0 by auto }
then show ?thesis by (simp add: Limsup_const)
next
- assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
+ assume neq_0: "(\<integral>\<^sup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
have "0 = limsup (\<lambda>n. 0 :: ereal)" by (simp add: Limsup_const)
- also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
+ also have "\<dots> \<le> limsup (\<lambda>n. \<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
by (simp add: Limsup_mono positive_integral_positive)
- finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
- have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
+ finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
+ have "?wx = (\<integral>\<^sup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
using u'
proof (intro positive_integral_cong_AE, eventually_elim)
fix x assume u': "(\<lambda>i. u i x) ----> u' x"
@@ -2162,18 +2162,18 @@
by (auto intro!: tendsto_real_max)
qed (rule trivial_limit_sequentially)
qed
- also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
+ also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^sup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
using w u unfolding integrable_def
by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
- also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
- limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. ereal (2 * w x) \<partial>M) -
+ limsup (\<lambda>n. \<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
unfolding PI_diff positive_integral_max_0
using positive_integral_positive[of M "\<lambda>x. ereal (2 * w x)"]
by (subst liminf_ereal_cminus) auto
finally show ?thesis
using neq_0 I2w_fin positive_integral_positive[of M "\<lambda>x. ereal (2 * w x)"] pos
unfolding positive_integral_max_0
- by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"])
+ by (cases rule: ereal2_cases[of "\<integral>\<^sup>+ x. ereal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"])
auto
qed
@@ -2186,7 +2186,7 @@
finally have "0 \<le> liminf ?f" . }
ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
using `limsup ?f = 0` by auto
- have "\<And>n. (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = ereal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
+ have "\<And>n. (\<integral>\<^sup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = ereal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
using diff positive_integral_positive[of M]
by (subst integral_eq_positive_integral[of _ M]) (auto simp: ereal_real integrable_def)
then show ?lim_diff
@@ -2200,15 +2200,15 @@
obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M) < r"
using diff by (auto simp: integral_positive)
- show "\<exists>N. \<forall>n\<ge>N. norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r"
+ show "\<exists>N. \<forall>n\<ge>N. norm (integral\<^sup>L M (u n) - integral\<^sup>L M u') < r"
proof (safe intro!: exI[of _ N])
fix n assume "N \<le> n"
- have "\<bar>integral\<^isup>L M (u n) - integral\<^isup>L M u'\<bar> = \<bar>(\<integral>x. u n x - u' x \<partial>M)\<bar>"
+ have "\<bar>integral\<^sup>L M (u n) - integral\<^sup>L M u'\<bar> = \<bar>(\<integral>x. u n x - u' x \<partial>M)\<bar>"
using u `integrable M u'` by auto
also have "\<dots> \<le> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)" using u `integrable M u'`
by (rule_tac integral_triangle_inequality) auto
also note N[OF `N \<le> n`]
- finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp
+ finally show "norm (integral\<^sup>L M (u n) - integral\<^sup>L M u') < r" by simp
qed
qed
qed
@@ -2218,7 +2218,7 @@
and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
and sums: "summable (\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
shows "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")
- and "(\<lambda>i. integral\<^isup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is ?integral)
+ and "(\<lambda>i. integral\<^sup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is ?integral)
proof -
have "\<forall>x\<in>space M. \<exists>w. (\<lambda>i. \<bar>f i x\<bar>) sums w"
using summable unfolding summable_def by auto
@@ -2255,10 +2255,10 @@
by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
show "AE x in M. (\<lambda>n. ?w' n x) ----> ?w x"
using w by (simp_all add: tendsto_const sums_def)
- have *: "\<And>n. integral\<^isup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
+ have *: "\<And>n. integral\<^sup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
using integrable by (simp add: integrable_abs cong: integral_cong)
from abs_sum
- show "(\<lambda>i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def .
+ show "(\<lambda>i. integral\<^sup>L M (?w' i)) ----> x" unfolding * sums_def .
qed (simp add: w_borel measurable_If_set)
from summable[THEN summable_rabs_cancel]
@@ -2295,7 +2295,7 @@
using f(2) by (intro AE_I2) (auto split: split_indicator)
have int: "\<And>n. integrable M (\<lambda>x. f x * indicator {.. n} x)"
by (rule integrable_mult_indicator) (auto simp: M f)
- show "(\<lambda>n. \<integral> x. f x * indicator {..real n} x \<partial>M) ----> integral\<^isup>L M f"
+ show "(\<lambda>n. \<integral> x. f x * indicator {..real n} x \<partial>M) ----> integral\<^sup>L M f"
proof (rule integral_dominated_convergence)
{ fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially"
by (rule eventually_sequentiallyI[of "natceiling x"])
@@ -2310,8 +2310,8 @@
note nonneg = this
let ?P = "\<lambda>y. \<integral> x. max 0 (f x) * indicator {..y} x \<partial>M"
let ?N = "\<lambda>y. \<integral> x. max 0 (- f x) * indicator {..y} x \<partial>M"
- let ?p = "integral\<^isup>L M (\<lambda>x. max 0 (f x))"
- let ?n = "integral\<^isup>L M (\<lambda>x. max 0 (- f x))"
+ let ?p = "integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+ let ?n = "integral\<^sup>L M (\<lambda>x. max 0 (- f x))"
have "(?P ---> ?p) at_top" "(?N ---> ?n) at_top"
by (auto intro!: nonneg integrable_max f)
note tendsto_diff[OF this]
@@ -2319,7 +2319,7 @@
by (subst integral_diff(2)[symmetric])
(auto intro!: integrable_mult_indicator integrable_max f integral_cong ext
simp: M split: split_max)
- also have "?p - ?n = integral\<^isup>L M f"
+ also have "?p - ?n = integral\<^sup>L M f"
by (subst integral_diff(2)[symmetric])
(auto intro!: integrable_max f integral_cong ext simp: M split: split_max)
finally show ?thesis .
@@ -2332,7 +2332,7 @@
assumes borel: "f \<in> borel_measurable borel"
assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)"
assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) ---> x) at_top"
- shows "integrable M f" "integral\<^isup>L M f = x"
+ shows "integrable M f" "integral\<^sup>L M f = x"
proof -
from nonneg have "AE x in M. mono (\<lambda>n::nat. f x * indicator {..real n} x)"
by (auto split: split_indicator intro!: monoI)
@@ -2350,7 +2350,7 @@
using borel by simp
show "integrable M f"
by (rule integral_monotone_convergence) fact+
- show "integral\<^isup>L M f = x"
+ show "integral\<^sup>L M f = x"
by (rule integral_monotone_convergence) fact+
qed
@@ -2364,11 +2364,11 @@
and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> (emeasure M) (f -` {x} \<inter> space M) \<noteq> \<infinity>"
and abs_summable: "summable (\<lambda>r. \<bar>enum r * real ((emeasure M) (f -` {enum r} \<inter> space M))\<bar>)"
shows "integrable M f"
- and "(\<lambda>r. enum r * real ((emeasure M) (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
+ and "(\<lambda>r. enum r * real ((emeasure M) (f -` {enum r} \<inter> space M))) sums integral\<^sup>L M f" (is ?sums)
proof -
let ?A = "\<lambda>r. f -` {enum r} \<inter> space M"
let ?F = "\<lambda>r x. enum r * indicator (?A r) x"
- have enum_eq: "\<And>r. enum r * real ((emeasure M) (?A r)) = integral\<^isup>L M (?F r)"
+ have enum_eq: "\<And>r. enum r * real ((emeasure M) (?A r)) = integral\<^sup>L M (?F r)"
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
{ fix x assume "x \<in> space M"
@@ -2389,7 +2389,7 @@
by (auto intro!: sums_single simp: F F_abs) }
note F_sums_f = this(1) and F_abs_sums_f = this(2)
- have int_f: "integral\<^isup>L M f = (\<integral>x. (\<Sum>r. ?F r x) \<partial>M)" "integrable M f = integrable M (\<lambda>x. \<Sum>r. ?F r x)"
+ have int_f: "integral\<^sup>L M f = (\<integral>x. (\<Sum>r. ?F r x) \<partial>M)" "integrable M f = integrable M (\<lambda>x. \<Sum>r. ?F r x)"
using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff)
{ fix r
@@ -2419,7 +2419,7 @@
lemma positive_integral_distr':
assumes T: "T \<in> measurable M M'"
and f: "f \<in> borel_measurable (distr M M' T)" "\<And>x. 0 \<le> f x"
- shows "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
+ shows "integral\<^sup>P (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
using f
proof induct
case (cong f g)
@@ -2441,12 +2441,12 @@
positive_integral_monotone_convergence_SUP le_fun_def incseq_def)
lemma positive_integral_distr:
- "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
+ "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^sup>P (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
by (subst (1 2) positive_integral_max_0[symmetric])
(simp add: positive_integral_distr')
lemma integral_distr:
- "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^isup>L (distr M M' T) f = (\<integral> x. f (T x) \<partial>M)"
+ "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^sup>L (distr M M' T) f = (\<integral> x. f (T x) \<partial>M)"
unfolding lebesgue_integral_def
by (subst (1 2) positive_integral_distr) auto
@@ -2467,13 +2467,13 @@
lemma positive_integral_count_space:
assumes A: "finite {a\<in>A. 0 < f a}"
- shows "integral\<^isup>P (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
+ shows "integral\<^sup>P (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
proof -
- have *: "(\<integral>\<^isup>+x. max 0 (f x) \<partial>count_space A) =
- (\<integral>\<^isup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
+ have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
+ (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
by (auto intro!: positive_integral_cong
simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less)
- also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^isup>+ x. f a * indicator {a} x \<partial>count_space A)"
+ also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
by (subst positive_integral_setsum)
(simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le)
also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
@@ -2486,7 +2486,7 @@
by (auto simp: positive_integral_count_space integrable_def)
lemma positive_integral_count_space_finite:
- "finite A \<Longrightarrow> (\<integral>\<^isup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))"
+ "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))"
by (subst positive_integral_max_0[symmetric])
(auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le)
@@ -2525,7 +2525,7 @@
section {* Measure spaces with an associated density *}
definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
- "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
+ "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
lemma
shows sets_density[simp]: "sets (density M f) = sets M"
@@ -2559,14 +2559,14 @@
lemma emeasure_density:
assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M"
- shows "emeasure (density M f) A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
+ shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
(is "_ = ?\<mu> A")
unfolding density_def
proof (rule emeasure_measure_of_sigma)
show "sigma_algebra (space M) (sets M)" ..
show "positive (sets M) ?\<mu>"
using f by (auto simp: positive_def intro!: positive_integral_positive)
- have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^isup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'")
+ have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^sup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'")
apply (subst positive_integral_max_0[symmetric])
apply (intro ext positive_integral_cong_AE AE_I2)
apply (auto simp: indicator_def)
@@ -2579,11 +2579,11 @@
then have *: "\<And>i. (\<lambda>x. max 0 (f x) * indicator (A i) x) \<in> borel_measurable M"
by (auto simp: set_eq_iff)
assume disj: "disjoint_family A"
- have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)"
+ have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)"
using f * by (simp add: positive_integral_suminf)
- also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f
+ also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f
by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE)
- also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)"
unfolding suminf_indicator[OF disj] ..
finally show "(\<Sum>n. ?\<mu>' (A n)) = ?\<mu>' (\<Union>x. A x)" by simp
qed
@@ -2594,12 +2594,12 @@
shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
proof -
{ assume "A \<in> sets M"
- have eq: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. max 0 (f x) * indicator A x \<partial>M)"
+ have eq: "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. max 0 (f x) * indicator A x \<partial>M)"
apply (subst positive_integral_max_0[symmetric])
apply (intro positive_integral_cong)
apply (auto simp: indicator_def)
done
- have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow>
+ have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow>
emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
unfolding eq
using f `A \<in> sets M`
@@ -2609,7 +2609,7 @@
by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
also have "(AE x in M. max 0 (f x) * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
by (auto simp add: indicator_def max_def split: split_if_asm)
- finally have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
+ finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
with f show ?thesis
by (simp add: null_sets_def emeasure_density cong: conj_cong)
qed
@@ -2641,7 +2641,7 @@
lemma positive_integral_density':
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
assumes g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
- shows "integral\<^isup>P (density M f) g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
+ shows "integral\<^sup>P (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
using g proof induct
case (cong u v)
then show ?case
@@ -2677,7 +2677,7 @@
lemma positive_integral_density:
"f \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> g' \<in> borel_measurable M \<Longrightarrow>
- integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+ x. f x * g' x \<partial>M)"
+ integral\<^sup>P (density M f) g' = (\<integral>\<^sup>+ x. f x * g' x \<partial>M)"
by (subst (1 2) positive_integral_max_0[symmetric])
(auto intro!: positive_integral_cong_AE
simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density')
@@ -2685,7 +2685,7 @@
lemma integral_density:
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
and g: "g \<in> borel_measurable M"
- shows "integral\<^isup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
+ shows "integral\<^sup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
and "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>x. f x * g x)"
unfolding lebesgue_integral_def integrable_def using f g
by (auto simp: positive_integral_density)
@@ -2694,9 +2694,9 @@
assumes S: "S \<in> sets M" and X: "X \<in> sets M"
shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)"
proof -
- have "emeasure (density M (indicator S)) X = (\<integral>\<^isup>+x. indicator S x * indicator X x \<partial>M)"
+ have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)"
using S X by (simp add: emeasure_density)
- also have "\<dots> = (\<integral>\<^isup>+x. indicator (S \<inter> X) x \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)"
by (auto intro!: positive_integral_cong simp: indicator_def)
also have "\<dots> = emeasure M (S \<inter> X)"
using S X by (simp add: sets.Int)
@@ -2810,7 +2810,7 @@
lemma positive_integral_point_measure:
"finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
- integral\<^isup>P (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
+ integral\<^sup>P (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
unfolding point_measure_def
apply (subst density_max_0)
apply (subst positive_integral_density)
@@ -2825,11 +2825,11 @@
lemma positive_integral_point_measure_finite:
"finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> g a) \<Longrightarrow>
- integral\<^isup>P (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+ integral\<^sup>P (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le)
lemma lebesgue_integral_point_measure_finite:
- "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> integral\<^isup>L (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+ "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> integral\<^sup>L (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def)
lemma integrable_point_measure_finite:
@@ -2853,7 +2853,7 @@
assumes A: "A \<in> sets M" and B: "B \<in> sets M"
shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
proof -
- from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^isup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
+ from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
by (auto simp add: uniform_measure_def emeasure_density split: split_indicator
intro!: positive_integral_cong)
also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
--- a/src/HOL/Probability/Lebesgue_Measure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -513,7 +513,7 @@
assumes sets_eq: "sets M = sets borel"
shows "lborel = M"
proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
- let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
+ let ?P = "\<Pi>\<^sub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
by (simp_all add: borel_eq_atLeastAtMost sets_eq)
@@ -564,7 +564,7 @@
assumes f:"simple_function lebesgue f"
and f':"range f \<subseteq> {0..<\<infinity>}"
and om:"\<And>x. x \<in> range f \<Longrightarrow> emeasure lebesgue (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
- shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
+ shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^sup>S lebesgue f))) UNIV"
unfolding simple_integral_def space_lebesgue
proof (subst lebesgue_simple_function_indicator)
let ?M = "\<lambda>x. emeasure lebesgue (f -` {x} \<inter> UNIV)"
@@ -600,8 +600,8 @@
lemma simple_function_has_integral':
fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"
- and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"
- shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
+ and i: "integral\<^sup>S lebesgue f \<noteq> \<infinity>"
+ shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^sup>S lebesgue f))) UNIV"
proof -
let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
note f(1)[THEN simple_functionD(2)]
@@ -612,7 +612,7 @@
have "AE x in lebesgue. f x = ?f x"
using simple_integral_PInf[OF f i]
by (intro AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
- from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"
+ from f(1) f' this have eq: "integral\<^sup>S lebesgue f = integral\<^sup>S lebesgue ?f"
by (rule simple_integral_cong_AE)
have real_eq: "\<And>x. real (f x) = real (?f x)" by auto
@@ -620,10 +620,10 @@
unfolding eq real_eq
proof (rule simple_function_has_integral[OF f' rng])
fix x assume x: "x \<in> range ?f" and inf: "emeasure lebesgue (?f -` {x} \<inter> UNIV) = \<infinity>"
- have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
+ have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^sup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
using f'[THEN simple_functionD(2)]
by (simp add: simple_integral_cmult_indicator)
- also have "\<dots> \<le> integral\<^isup>S lebesgue f"
+ also have "\<dots> \<le> integral\<^sup>S lebesgue f"
using f'[THEN simple_functionD(2)] f
by (intro simple_integral_mono simple_function_mult simple_function_indicator)
(auto split: split_indicator)
@@ -633,8 +633,8 @@
lemma positive_integral_has_integral:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
- assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
- shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
+ assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^sup>P lebesgue f \<noteq> \<infinity>"
+ shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^sup>P lebesgue f))) UNIV"
proof -
from borel_measurable_implies_simple_function_sequence'[OF f(1)]
guess u . note u = this
@@ -644,12 +644,12 @@
note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric]
{ fix i
note u_eq
- also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"
+ also have "integral\<^sup>P lebesgue (u i) \<le> (\<integral>\<^sup>+x. max 0 (f x) \<partial>lebesgue)"
by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric])
- finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"
+ finally have "integral\<^sup>S lebesgue (u i) \<noteq> \<infinity>"
unfolding positive_integral_max_0 using f by auto }
note u_fin = this
- then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"
+ then have u_int: "\<And>i. (?u i has_integral real (integral\<^sup>S lebesgue (u i))) UNIV"
by (rule simple_function_has_integral'[OF u(1,5)])
have "\<forall>x. \<exists>r\<ge>0. f x = ereal r"
proof
@@ -684,33 +684,33 @@
fix k
have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)"
by (intro abs_of_nonneg integral_nonneg int ballI u')
- also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"
+ also have "\<dots> = real (integral\<^sup>S lebesgue (u k))"
using u_int[THEN integral_unique] by (simp add: u')
- also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
+ also have "\<dots> = real (integral\<^sup>P lebesgue (u k))"
using positive_integral_eq_simple_integral[OF u(1,5)] by simp
- also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
+ also have "\<dots> \<le> real (integral\<^sup>P lebesgue f)" using f
by (auto intro!: real_of_ereal_positive_mono positive_integral_positive
positive_integral_mono SUP_upper simp: SUP_eq[symmetric])
- finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
+ finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^sup>P lebesgue f)" .
qed
qed
- have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
+ have "integral\<^sup>P lebesgue f = ereal (integral UNIV f')"
proof (rule tendsto_unique[OF trivial_limit_sequentially])
- have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
+ have "(\<lambda>i. integral\<^sup>S lebesgue (u i)) ----> (SUP i. integral\<^sup>P lebesgue (u i))"
unfolding u_eq by (intro LIMSEQ_SUP incseq_positive_integral u)
also note positive_integral_monotone_convergence_SUP
[OF u(2) borel_measurable_simple_function[OF u(1)] u(5), symmetric]
- finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
+ finally show "(\<lambda>k. integral\<^sup>S lebesgue (u k)) ----> integral\<^sup>P lebesgue f"
unfolding SUP_eq .
{ fix k
- have "0 \<le> integral\<^isup>S lebesgue (u k)"
+ have "0 \<le> integral\<^sup>S lebesgue (u k)"
using u by (auto intro!: simple_integral_positive)
- then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"
+ then have "integral\<^sup>S lebesgue (u k) = ereal (real (integral\<^sup>S lebesgue (u k)))"
using u_fin by (auto simp: ereal_real) }
note * = this
- show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
+ show "(\<lambda>k. integral\<^sup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
using convergent using u_int[THEN integral_unique, symmetric]
by (subst *) (simp add: u')
qed
@@ -720,11 +720,11 @@
lemma lebesgue_integral_has_integral:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "integrable lebesgue f"
- shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"
+ shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
proof -
let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"
have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)
- { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
+ { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^sup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
by (intro positive_integral_cong_pos) (auto split: split_max) }
note eq = this
show ?thesis
@@ -740,16 +740,16 @@
lemma lebesgue_simple_integral_eq_borel:
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>S lebesgue f = integral\<^isup>S lborel f"
+ shows "integral\<^sup>S lebesgue f = integral\<^sup>S lborel f"
using f[THEN measurable_sets]
by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_lborel[symmetric]
simp: simple_integral_def)
lemma lebesgue_positive_integral_eq_borel:
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
+ shows "integral\<^sup>P lebesgue f = integral\<^sup>P lborel f"
proof -
- from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"
+ from f have "integral\<^sup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^sup>P lborel (\<lambda>x. max 0 (f x))"
by (auto intro!: positive_integral_subalgebra[symmetric])
then show ?thesis unfolding positive_integral_max_0 .
qed
@@ -757,7 +757,7 @@
lemma lebesgue_integral_eq_borel:
assumes "f \<in> borel_measurable borel"
shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P)
- and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I)
+ and "integral\<^sup>L lebesgue f = integral\<^sup>L lborel f" (is ?I)
proof -
have "sets lborel \<subseteq> sets lebesgue" by auto
from integral_subalgebra[of f lborel, OF _ this _ _] assms
@@ -767,7 +767,7 @@
lemma borel_integral_has_integral:
fixes f::"'a::ordered_euclidean_space => real"
assumes f:"integrable lborel f"
- shows "(f has_integral (integral\<^isup>L lborel f)) UNIV"
+ shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
proof -
have borel: "f \<in> borel_measurable borel"
using f unfolding integrable_def by auto
@@ -780,12 +780,12 @@
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f_borel: "f \<in> borel_measurable lebesgue" and nonneg: "\<And>x. 0 \<le> f x"
assumes I: "(f has_integral I) UNIV"
- shows "(\<integral>\<^isup>+x. f x \<partial>lebesgue) = I"
+ shows "(\<integral>\<^sup>+x. f x \<partial>lebesgue) = I"
proof -
from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue" by auto
from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
- have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^isup>S lebesgue (F i))"
+ have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^sup>S lebesgue (F i))"
using F
by (subst positive_integral_monotone_convergence_simple)
(simp_all add: positive_integral_max_0 simple_function_def)
@@ -835,7 +835,7 @@
by (simp add: integrable_on_cmult_iff) }
note F_finite = lmeasure_finite[OF this]
- have "((\<lambda>x. real (F i x)) has_integral real (integral\<^isup>S lebesgue (F i))) UNIV"
+ have "((\<lambda>x. real (F i x)) has_integral real (integral\<^sup>S lebesgue (F i))) UNIV"
proof (rule simple_function_has_integral[of "F i"])
show "simple_function lebesgue (F i)"
using F(1) by (simp add: simple_function_def)
@@ -845,7 +845,7 @@
fix x assume "x \<in> range (F i)" "emeasure lebesgue (F i -` {x} \<inter> UNIV) = \<infinity>"
with F_finite[of x] show "x = 0" by auto
qed
- from this I have "real (integral\<^isup>S lebesgue (F i)) \<le> I"
+ from this I have "real (integral\<^sup>S lebesgue (F i)) \<le> I"
by (rule has_integral_le) (intro ballI F_bound)
moreover
{ fix x assume x: "x \<in> range (F i)"
@@ -853,37 +853,37 @@
by (auto simp: image_iff le_less) metis
with F_finite[OF _ x] x have "x * emeasure lebesgue (F i -` {x} \<inter> UNIV) \<noteq> \<infinity>"
by auto }
- then have "integral\<^isup>S lebesgue (F i) \<noteq> \<infinity>"
+ then have "integral\<^sup>S lebesgue (F i) \<noteq> \<infinity>"
unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast
- moreover have "0 \<le> integral\<^isup>S lebesgue (F i)"
+ moreover have "0 \<le> integral\<^sup>S lebesgue (F i)"
using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def)
- ultimately show "integral\<^isup>S lebesgue (F i) \<le> ereal I"
- by (cases "integral\<^isup>S lebesgue (F i)") auto
+ ultimately show "integral\<^sup>S lebesgue (F i) \<le> ereal I"
+ by (cases "integral\<^sup>S lebesgue (F i)") auto
qed
also have "\<dots> < \<infinity>" by simp
- finally have finite: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp
+ finally have finite: "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp
have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue"
using f_borel by (auto intro: borel_measurable_lebesgueI)
from positive_integral_has_integral[OF borel _ finite]
- have "(f has_integral real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)) UNIV"
+ have "(f has_integral real (\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue)) UNIV"
using nonneg by (simp add: subset_eq)
- with I have "I = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)"
+ with I have "I = real (\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue)"
by (rule has_integral_unique)
with finite positive_integral_positive[of _ "\<lambda>x. ereal (f x)"] show ?thesis
- by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue") auto
+ by (cases "\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue") auto
qed
lemma has_integral_iff_positive_integral_lebesgue:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x"
- shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lebesgue f = I"
+ shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^sup>P lebesgue f = I"
using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f]
by (auto simp: subset_eq)
lemma has_integral_iff_positive_integral_lborel:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x"
- shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lborel f = I"
+ shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^sup>P lborel f = I"
using assms
by (subst has_integral_iff_positive_integral_lebesgue)
(auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel)
@@ -912,17 +912,17 @@
lemma sets_product_borel:
assumes I: "finite I"
- shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
+ shows "sets (\<Pi>\<^sub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^sub>E i\<in>I. UNIV) { \<Pi>\<^sub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
- show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
+ show "sigma_sets (space (Pi\<^sub>M I (\<lambda>i. lborel))) {Pi\<^sub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
qed (auto simp: borel_eq_lessThan reals_Archimedean2)
lemma measurable_e2p[measurable]:
- "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))"
+ "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure))"
proof (rule measurable_sigma_sets[OF sets_product_borel])
- fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i} |x. True} "
- then obtain x where "A = (\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i})" by auto
+ fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i} |x. True} "
+ then obtain x where "A = (\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i})" by auto
then have "e2p -` A = {..< (\<Sum>i\<in>Basis. x i *\<^sub>R i) :: 'a}"
using DIM_positive by (auto simp add: set_eq_iff e2p_def
euclidean_eq_iff[where 'a='a] eucl_less[where 'a='a])
@@ -934,27 +934,27 @@
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
lemma measurable_p2e[measurable]:
- "p2e \<in> measurable (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))
+ "p2e \<in> measurable (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure))
(borel :: 'a::ordered_euclidean_space measure)"
(is "p2e \<in> measurable ?P _")
proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
fix x and i :: 'a
let ?A = "{w \<in> space ?P. (p2e w :: 'a) \<bullet> i \<le> x}"
assume "i \<in> Basis"
- then have "?A = (\<Pi>\<^isub>E j\<in>Basis. if i = j then {.. x} else UNIV)"
+ then have "?A = (\<Pi>\<^sub>E j\<in>Basis. if i = j then {.. x} else UNIV)"
using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm)
then show "?A \<in> sets ?P"
by auto
qed
lemma lborel_eq_lborel_space:
- "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e"
+ "(lborel :: 'a measure) = distr (\<Pi>\<^sub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e"
(is "?B = ?D")
proof (rule lborel_eqI)
show "sets ?D = sets borel" by simp
- let ?P = "(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"
+ let ?P = "(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)"
fix a b :: 'a
- have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>Basis. {a \<bullet> i .. b \<bullet> i})"
+ have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^sub>E i\<in>Basis. {a \<bullet> i .. b \<bullet> i})"
by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff)
have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
proof cases
@@ -975,12 +975,12 @@
lemma borel_fubini_positiv_integral:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"
+ shows "integral\<^sup>P lborel f = \<integral>\<^sup>+x. f (p2e x) \<partial>(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)"
by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
lemma borel_fubini_integrable:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
- shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))"
+ shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))"
(is "_ \<longleftrightarrow> integrable ?B ?f")
proof
assume "integrable lborel f"
@@ -1005,7 +1005,7 @@
lemma borel_fubini:
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
assumes f: "f \<in> borel_measurable borel"
- shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel))"
+ shows "integral\<^sup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel))"
using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
lemma integrable_on_borel_integrable:
@@ -1014,10 +1014,10 @@
assumes f: "f integrable_on UNIV"
shows "integrable lborel f"
proof -
- have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel) \<noteq> \<infinity>"
+ have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lborel) \<noteq> \<infinity>"
using has_integral_iff_positive_integral_lborel[OF f_borel nonneg] f
by (auto simp: integrable_on_def)
- moreover have "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>lborel) = 0"
+ moreover have "(\<integral>\<^sup>+ x. ereal (- f x) \<partial>lborel) = 0"
using f_borel nonneg by (subst positive_integral_0_iff_AE) auto
ultimately show ?thesis
using f_borel by (auto simp: integrable_def)
@@ -1065,7 +1065,7 @@
assumes "a \<le> b"
and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
and f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
- shows "integral\<^isup>L lborel (\<lambda>x. f x * indicator {a .. b} x) = F b - F a"
+ shows "integral\<^sup>L lborel (\<lambda>x. f x * indicator {a .. b} x) = F b - F a"
proof -
let ?f = "\<lambda>x. f x * indicator {a .. b} x"
have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
@@ -1078,7 +1078,7 @@
by (subst has_integral_eq_eq[where g=f]) auto
then have "(?f has_integral F b - F a) UNIV"
by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto
- ultimately show "integral\<^isup>L lborel ?f = F b - F a"
+ ultimately show "integral\<^sup>L lborel ?f = F b - F a"
by (rule has_integral_unique)
qed
@@ -1091,7 +1091,7 @@
lemma positive_integral_FTC_atLeastAtMost:
assumes f_borel: "f \<in> borel_measurable borel"
assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
- shows "(\<integral>\<^isup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
+ shows "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
proof -
have i: "(f has_integral F b - F a) {a..b}"
by (intro fundamental_theorem_of_calculus ballI has_vector_derivative_withinI_DERIV assms)
@@ -1099,10 +1099,10 @@
by (rule has_integral_eq[OF _ i]) auto
have i: "((\<lambda>x. f x * indicator {a..b} x) has_integral F b - F a) UNIV"
by (rule has_integral_on_superset[OF _ _ i]) auto
- then have "(\<integral>\<^isup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = F b - F a"
+ then have "(\<integral>\<^sup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = F b - F a"
using f f_borel
by (subst has_integral_iff_positive_integral_lborel[symmetric]) (auto split: split_indicator)
- also have "(\<integral>\<^isup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = (\<integral>\<^isup>+x. ereal (f x) * indicator {a .. b} x \<partial>lborel)"
+ also have "(\<integral>\<^sup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = (\<integral>\<^sup>+x. ereal (f x) * indicator {a .. b} x \<partial>lborel)"
by (auto intro!: positive_integral_cong simp: indicator_def)
finally show ?thesis by simp
qed
@@ -1113,7 +1113,7 @@
assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
assumes lim: "(F ---> T) at_top"
- shows "(\<integral>\<^isup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
+ shows "(\<integral>\<^sup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
proof -
let ?f = "\<lambda>(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x"
let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x"
@@ -1129,9 +1129,9 @@
then show "(\<lambda>n. ?f n x) ----> ?fR x"
by (rule Lim_eventually)
qed
- then have "integral\<^isup>P lborel ?fR = (\<integral>\<^isup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
+ then have "integral\<^sup>P lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
by simp
- also have "\<dots> = (SUP i::nat. (\<integral>\<^isup>+ x. ?f i x \<partial>lborel))"
+ also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
proof (rule positive_integral_monotone_convergence_SUP)
show "incseq ?f"
using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
--- a/src/HOL/Probability/Probability_Measure.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy Tue Aug 13 16:25:47 2013 +0200
@@ -26,7 +26,7 @@
abbreviation (in prob_space) "events \<equiv> sets M"
abbreviation (in prob_space) "prob \<equiv> measure M"
abbreviation (in prob_space) "random_variable M' X \<equiv> X \<in> measurable M M'"
-abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
+abbreviation (in prob_space) "expectation \<equiv> integral\<^sup>L M"
lemma (in prob_space) prob_space_distr:
assumes f: "f \<in> measurable M M'" shows "prob_space (distr M M' f)"
@@ -278,19 +278,19 @@
lemma (in prob_space) joint_distribution_Times_le_fst:
"random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
- \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MX X) A"
+ \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MX X) A"
by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
lemma (in prob_space) joint_distribution_Times_le_snd:
"random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
- \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MY Y) B"
+ \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^sub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MY Y) B"
by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2
-sublocale pair_prob_space \<subseteq> P: prob_space "M1 \<Otimes>\<^isub>M M2"
+sublocale pair_prob_space \<subseteq> P: prob_space "M1 \<Otimes>\<^sub>M M2"
proof
- show "emeasure (M1 \<Otimes>\<^isub>M M2) (space (M1 \<Otimes>\<^isub>M M2)) = 1"
+ show "emeasure (M1 \<Otimes>\<^sub>M M2) (space (M1 \<Otimes>\<^sub>M M2)) = 1"
by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure)
qed
@@ -303,17 +303,17 @@
locale finite_product_prob_space = finite_product_sigma_finite M I + product_prob_space M I for M I
-sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^isub>M i\<in>I. M i"
+sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^sub>M i\<in>I. M i"
proof
- show "emeasure (\<Pi>\<^isub>M i\<in>I. M i) (space (\<Pi>\<^isub>M i\<in>I. M i)) = 1"
+ show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (space (\<Pi>\<^sub>M i\<in>I. M i)) = 1"
by (simp add: measure_times M.emeasure_space_1 setprod_1 space_PiM)
qed
lemma (in finite_product_prob_space) prob_times:
assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
- shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
+ shows "prob (\<Pi>\<^sub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
proof -
- have "ereal (measure (\<Pi>\<^isub>M i\<in>I. M i) (\<Pi>\<^isub>E i\<in>I. X i)) = emeasure (\<Pi>\<^isub>M i\<in>I. M i) (\<Pi>\<^isub>E i\<in>I. X i)"
+ have "ereal (measure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)) = emeasure (\<Pi>\<^sub>M i\<in>I. M i) (\<Pi>\<^sub>E i\<in>I. X i)"
using X by (simp add: emeasure_eq_measure)
also have "\<dots> = (\<Prod>i\<in>I. emeasure (M i) (X i))"
using measure_times X by simp
@@ -357,7 +357,7 @@
by simp_all
lemma
- assumes D: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) f"
+ assumes D: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f"
shows joint_distributed_measurable1[measurable_dest]:
"h1 \<in> measurable N M \<Longrightarrow> (\<lambda>x. X (h1 x)) \<in> measurable N S"
and joint_distributed_measurable2[measurable_dest]:
@@ -374,7 +374,7 @@
using X a A by (simp add: emeasure_distr)
also have "\<dots> = emeasure (density (count_space A) P) {a}"
using X by (simp add: distributed_distr_eq_density)
- also have "\<dots> = (\<integral>\<^isup>+x. P a * indicator {a} x \<partial>count_space A)"
+ also have "\<dots> = (\<integral>\<^sup>+x. P a * indicator {a} x \<partial>count_space A)"
using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: positive_integral_cong)
also have "\<dots> = P a"
using X a by (subst positive_integral_cmult_indicator) (auto simp: distributed_def one_ereal_def[symmetric] AE_count_space)
@@ -415,12 +415,12 @@
using subdensity[OF T, of M X "\<lambda>x. ereal (f x)" Y "\<lambda>x. ereal (g x)"] assms by auto
lemma distributed_emeasure:
- "distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^isup>+x. f x * indicator A x \<partial>N)"
+ "distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>N)"
by (auto simp: distributed_AE
distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr)
lemma distributed_positive_integral:
- "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^isup>+x. f x * g x \<partial>N) = (\<integral>\<^isup>+x. g (X x) \<partial>M)"
+ "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f x * g x \<partial>N) = (\<integral>\<^sup>+x. g (X x) \<partial>M)"
by (auto simp: distributed_AE
distributed_distr_eq_density[symmetric] positive_integral_density[symmetric] positive_integral_distr)
@@ -460,10 +460,10 @@
lemma (in prob_space) distributed_jointI:
assumes "sigma_finite_measure S" "sigma_finite_measure T"
assumes X[measurable]: "X \<in> measurable M S" and Y[measurable]: "Y \<in> measurable M T"
- assumes [measurable]: "f \<in> borel_measurable (S \<Otimes>\<^isub>M T)" and f: "AE x in S \<Otimes>\<^isub>M T. 0 \<le> f x"
+ assumes [measurable]: "f \<in> borel_measurable (S \<Otimes>\<^sub>M T)" and f: "AE x in S \<Otimes>\<^sub>M T. 0 \<le> f x"
assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow>
- emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)"
- shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) f"
+ emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)"
+ shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) f"
unfolding distributed_def
proof safe
interpret S: sigma_finite_measure S by fact
@@ -472,7 +472,7 @@
from ST.sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('b \<times> 'c) set" .. note F = this
let ?E = "{a \<times> b |a b. a \<in> sets S \<and> b \<in> sets T}"
- let ?P = "S \<Otimes>\<^isub>M T"
+ let ?P = "S \<Otimes>\<^sub>M T"
show "distr M ?P (\<lambda>x. (X x, Y x)) = density ?P f" (is "?L = ?R")
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]])
show "?E \<subseteq> Pow (space ?P)"
@@ -492,7 +492,7 @@
and A[measurable]: "A \<in> sets S" and B[measurable]: "B \<in> sets T" by auto
have "emeasure ?L E = emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B}"
by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair)
- also have "\<dots> = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. (f (x, y) * indicator B y) * indicator A x \<partial>T) \<partial>S)"
+ also have "\<dots> = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. (f (x, y) * indicator B y) * indicator A x \<partial>T) \<partial>S)"
using f by (auto simp add: eq positive_integral_multc intro!: positive_integral_cong)
also have "\<dots> = emeasure ?R E"
by (auto simp add: emeasure_density T.positive_integral_fst_measurable(2)[symmetric]
@@ -503,8 +503,8 @@
lemma (in prob_space) distributed_swap:
assumes "sigma_finite_measure S" "sigma_finite_measure T"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- shows "distributed M (T \<Otimes>\<^isub>M S) (\<lambda>x. (Y x, X x)) (\<lambda>(x, y). Pxy (y, x))"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ shows "distributed M (T \<Otimes>\<^sub>M S) (\<lambda>x. (Y x, X x)) (\<lambda>(x, y). Pxy (y, x))"
proof -
interpret S: sigma_finite_measure S by fact
interpret T: sigma_finite_measure T by fact
@@ -516,27 +516,27 @@
apply (subst TS.distr_pair_swap)
unfolding distributed_def
proof safe
- let ?D = "distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x))"
+ let ?D = "distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))"
show 1: "(\<lambda>(x, y). Pxy (y, x)) \<in> borel_measurable ?D"
by auto
with Pxy
- show "AE x in distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x)). 0 \<le> (case x of (x, y) \<Rightarrow> Pxy (y, x))"
+ show "AE x in distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x)). 0 \<le> (case x of (x, y) \<Rightarrow> Pxy (y, x))"
by (subst AE_distr_iff)
(auto dest!: distributed_AE
simp: measurable_split_conv split_beta
intro!: measurable_Pair)
- show 2: "random_variable (distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))"
+ show 2: "random_variable (distr (S \<Otimes>\<^sub>M T) (T \<Otimes>\<^sub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))"
using Pxy by auto
- { fix A assume A: "A \<in> sets (T \<Otimes>\<^isub>M S)"
- let ?B = "(\<lambda>(x, y). (y, x)) -` A \<inter> space (S \<Otimes>\<^isub>M T)"
+ { fix A assume A: "A \<in> sets (T \<Otimes>\<^sub>M S)"
+ let ?B = "(\<lambda>(x, y). (y, x)) -` A \<inter> space (S \<Otimes>\<^sub>M T)"
from sets.sets_into_space[OF A]
have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
emeasure M ((\<lambda>x. (X x, Y x)) -` ?B \<inter> space M)"
by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure)
- also have "\<dots> = (\<integral>\<^isup>+ x. Pxy x * indicator ?B x \<partial>(S \<Otimes>\<^isub>M T))"
+ also have "\<dots> = (\<integral>\<^sup>+ x. Pxy x * indicator ?B x \<partial>(S \<Otimes>\<^sub>M T))"
using Pxy A by (intro distributed_emeasure) auto
finally have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
- (\<integral>\<^isup>+ x. Pxy x * indicator A (snd x, fst x) \<partial>(S \<Otimes>\<^isub>M T))"
+ (\<integral>\<^sup>+ x. Pxy x * indicator A (snd x, fst x) \<partial>(S \<Otimes>\<^sub>M T))"
by (auto intro!: positive_integral_cong split: split_indicator) }
note * = this
show "distr M ?D (\<lambda>x. (Y x, X x)) = density ?D (\<lambda>(x, y). Pxy (y, x))"
@@ -550,8 +550,8 @@
lemma (in prob_space) distr_marginal1:
assumes "sigma_finite_measure S" "sigma_finite_measure T"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- defines "Px \<equiv> \<lambda>x. (\<integral>\<^isup>+z. Pxy (x, z) \<partial>T)"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ defines "Px \<equiv> \<lambda>x. (\<integral>\<^sup>+z. Pxy (x, z) \<partial>T)"
shows "distributed M S X Px"
unfolding distributed_def
proof safe
@@ -565,9 +565,9 @@
show borel: "Px \<in> borel_measurable S"
by (auto intro!: T.positive_integral_fst_measurable simp: Px_def)
- interpret Pxy: prob_space "distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+ interpret Pxy: prob_space "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
by (intro prob_space_distr) simp
- have "(\<integral>\<^isup>+ x. max 0 (- Pxy x) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>\<^isup>+ x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
+ have "(\<integral>\<^sup>+ x. max 0 (- Pxy x) \<partial>(S \<Otimes>\<^sub>M T)) = (\<integral>\<^sup>+ x. 0 \<partial>(S \<Otimes>\<^sub>M T))"
using Pxy
by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_AE)
@@ -575,25 +575,25 @@
proof (rule measure_eqI)
fix A assume A: "A \<in> sets (distr M S X)"
with X measurable_space[of Y M T]
- have "emeasure (distr M S X) A = emeasure (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (A \<times> space T)"
+ have "emeasure (distr M S X) A = emeasure (distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))) (A \<times> space T)"
by (auto simp add: emeasure_distr intro!: arg_cong[where f="emeasure M"])
- also have "\<dots> = emeasure (density (S \<Otimes>\<^isub>M T) Pxy) (A \<times> space T)"
+ also have "\<dots> = emeasure (density (S \<Otimes>\<^sub>M T) Pxy) (A \<times> space T)"
using Pxy by (simp add: distributed_def)
- also have "\<dots> = \<integral>\<^isup>+ x. \<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
+ also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
using A borel Pxy
by (simp add: emeasure_density T.positive_integral_fst_measurable(2)[symmetric])
- also have "\<dots> = \<integral>\<^isup>+ x. Px x * indicator A x \<partial>S"
+ also have "\<dots> = \<integral>\<^sup>+ x. Px x * indicator A x \<partial>S"
apply (rule positive_integral_cong_AE)
using Pxy[THEN distributed_AE, THEN ST.AE_pair] AE_space
proof eventually_elim
fix x assume "x \<in> space S" "AE y in T. 0 \<le> Pxy (x, y)"
moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x"
by (auto simp: indicator_def)
- ultimately have "(\<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^isup>+ y. Pxy (x, y) \<partial>T) * indicator A x"
+ ultimately have "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) * indicator A x"
by (simp add: eq positive_integral_multc cong: positive_integral_cong)
- also have "(\<integral>\<^isup>+ y. Pxy (x, y) \<partial>T) = Px x"
+ also have "(\<integral>\<^sup>+ y. Pxy (x, y) \<partial>T) = Px x"
by (simp add: Px_def ereal_real positive_integral_positive)
- finally show "(\<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" .
+ finally show "(\<integral>\<^sup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" .
qed
finally show "emeasure (distr M S X) A = emeasure (density S Px) A"
using A borel Pxy by (simp add: emeasure_density)
@@ -605,31 +605,31 @@
lemma (in prob_space) distr_marginal2:
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- shows "distributed M T Y (\<lambda>y. (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S))"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ shows "distributed M T Y (\<lambda>y. (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S))"
using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp
lemma (in prob_space) distributed_marginal_eq_joint1:
assumes T: "sigma_finite_measure T"
assumes S: "sigma_finite_measure S"
assumes Px: "distributed M S X Px"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- shows "AE x in S. Px x = (\<integral>\<^isup>+y. Pxy (x, y) \<partial>T)"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ shows "AE x in S. Px x = (\<integral>\<^sup>+y. Pxy (x, y) \<partial>T)"
using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique)
lemma (in prob_space) distributed_marginal_eq_joint2:
assumes T: "sigma_finite_measure T"
assumes S: "sigma_finite_measure S"
assumes Py: "distributed M T Y Py"
- assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
- shows "AE y in T. Py y = (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S)"
+ assumes Pxy: "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+ shows "AE y in T. Py y = (\<integral>\<^sup>+x. Pxy (x, y) \<partial>S)"
using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique)
lemma (in prob_space) distributed_joint_indep':
assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
assumes X[measurable]: "distributed M S X Px" and Y[measurable]: "distributed M T Y Py"
- assumes indep: "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
- shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
+ assumes indep: "distr M S X \<Otimes>\<^sub>M distr M T Y = distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))"
+ shows "distributed M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
unfolding distributed_def
proof safe
interpret S: sigma_finite_measure S by fact
@@ -646,17 +646,17 @@
by (rule prob_space_distr) simp
have sf_Y: "sigma_finite_measure (density T Py)" ..
- show "distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). Px x * Py y)"
+ show "distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). Px x * Py y)"
unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y]
using distributed_borel_measurable[OF X] distributed_AE[OF X]
using distributed_borel_measurable[OF Y] distributed_AE[OF Y]
by (rule pair_measure_density[OF _ _ _ _ T sf_Y])
- show "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" by auto
+ show "random_variable (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x))" by auto
- show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^isub>M T)" by auto
+ show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^sub>M T)" by auto
- show "AE x in S \<Otimes>\<^isub>M T. 0 \<le> (case x of (x, y) \<Rightarrow> Px x * Py y)"
+ show "AE x in S \<Otimes>\<^sub>M T. 0 \<le> (case x of (x, y) \<Rightarrow> Px x * Py y)"
apply (intro ST.AE_pair_measure borel_measurable_le Pxy borel_measurable_const)
using distributed_AE[OF X]
apply eventually_elim
@@ -698,11 +698,11 @@
by (subst emeasure_distr)
(auto simp add: S_def P'_def simple_functionD emeasure_eq_measure measurable_count_space_eq2
intro!: arg_cong[where f=prob])
- also have "\<dots> = (\<integral>\<^isup>+x. ereal (P' a) * indicator {a} x \<partial>S)"
+ also have "\<dots> = (\<integral>\<^sup>+x. ereal (P' a) * indicator {a} x \<partial>S)"
using A X a
by (subst positive_integral_cmult_indicator)
(auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg)
- also have "\<dots> = (\<integral>\<^isup>+x. ereal (P' x) * indicator {a} x \<partial>S)"
+ also have "\<dots> = (\<integral>\<^sup>+x. ereal (P' x) * indicator {a} x \<partial>S)"
by (auto simp: indicator_def intro!: positive_integral_cong)
also have "\<dots> = emeasure (density S P') {a}"
using a A by (intro emeasure_density[symmetric]) (auto simp: S_def)
@@ -766,7 +766,7 @@
lemma (in prob_space) simple_distributed_joint:
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
- defines "S \<equiv> count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)"
+ defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M)"
defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then Px x else 0)"
shows "distributed M S (\<lambda>x. (X x, Y x)) P"
proof -
@@ -786,7 +786,7 @@
lemma (in prob_space) simple_distributed_joint2:
assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px"
- defines "S \<equiv> count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M) \<Otimes>\<^isub>M count_space (Z`space M)"
+ defines "S \<equiv> count_space (X`space M) \<Otimes>\<^sub>M count_space (Y`space M) \<Otimes>\<^sub>M count_space (Z`space M)"
defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then Px x else 0)"
shows "distributed M S (\<lambda>x. (X x, Y x, Z x)) P"
proof -
@@ -849,7 +849,7 @@
and A: "range A \<subseteq> E" "(\<Union>i::nat. A i) = space M1" "\<And>i. emeasure (distr M M1 X) (A i) \<noteq> \<infinity>"
and X: "X \<in> measurable M M1"
and f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
- and eq: "\<And>A. A \<in> E \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M1)"
+ and eq: "\<And>A. A \<in> E \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M1)"
shows "distributed M M1 X f"
unfolding distributed_def
proof (intro conjI)
@@ -880,7 +880,7 @@
fixes f :: "real \<Rightarrow> real"
assumes [measurable]: "X \<in> borel_measurable M"
and [measurable]: "f \<in> borel_measurable borel" and f[simp]: "AE x in lborel. 0 \<le> f x"
- and g_eq: "\<And>a. (\<integral>\<^isup>+x. f x * indicator {..a} x \<partial>lborel) = ereal (g a)"
+ and g_eq: "\<And>a. (\<integral>\<^sup>+x. f x * indicator {..a} x \<partial>lborel) = ereal (g a)"
and M_eq: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = ereal (g a)"
shows "distributed M lborel X f"
proof (rule distributedI_real)
@@ -896,7 +896,7 @@
fix A :: "real set" assume "A \<in> range atMost"
then obtain a where A: "A = {..a}" by auto
- show "emeasure M (X -` A \<inter> space M) = (\<integral>\<^isup>+x. f x * indicator A x \<partial>lborel)"
+ show "emeasure M (X -` A \<inter> space M) = (\<integral>\<^sup>+x. f x * indicator A x \<partial>lborel)"
unfolding vimage_eq A M_eq g_eq ..
qed auto
--- a/src/HOL/Probability/Projective_Family.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Projective_Family.thy Tue Aug 13 16:25:47 2013 +0200
@@ -11,27 +11,27 @@
lemma (in product_prob_space) distr_restrict:
assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
- shows "(\<Pi>\<^isub>M i\<in>J. M i) = distr (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D")
+ shows "(\<Pi>\<^sub>M i\<in>J. M i) = distr (\<Pi>\<^sub>M i\<in>K. M i) (\<Pi>\<^sub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D")
proof (rule measure_eqI_generator_eq)
have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset)
interpret J: finite_product_prob_space M J proof qed fact
interpret K: finite_product_prob_space M K proof qed fact
- let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
- let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)"
- let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
+ let ?J = "{Pi\<^sub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
+ let ?F = "\<lambda>i. \<Pi>\<^sub>E k\<in>J. space (M k)"
+ let ?\<Omega> = "(\<Pi>\<^sub>E k\<in>J. space (M k))"
show "Int_stable ?J"
by (rule Int_stable_PiE)
show "range ?F \<subseteq> ?J" "(\<Union>i. ?F i) = ?\<Omega>"
using `finite J` by (auto intro!: prod_algebraI_finite)
{ fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets.sets_into_space)
- show "sets (\<Pi>\<^isub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"
+ show "sets (\<Pi>\<^sub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"
using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
fix X assume "X \<in> ?J"
- then obtain E where [simp]: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
- with `finite J` have X: "X \<in> sets (Pi\<^isub>M J M)"
+ then obtain E where [simp]: "X = Pi\<^sub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
+ with `finite J` have X: "X \<in> sets (Pi\<^sub>M J M)"
by simp
have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))"
@@ -41,29 +41,29 @@
also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"
using `finite K` `J \<subseteq> K`
by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1)
- also have "\<dots> = emeasure (Pi\<^isub>M K M) (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))"
+ also have "\<dots> = emeasure (Pi\<^sub>M K M) (\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i))"
using E by (simp add: K.measure_times)
- also have "(\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i))"
+ also have "(\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^sub>E J E \<inter> (\<Pi>\<^sub>E i\<in>K. space (M i))"
using `J \<subseteq> K` sets.sets_into_space E by (force simp: Pi_iff PiE_def split: split_if_asm)
- finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X"
+ finally show "emeasure (Pi\<^sub>M J M) X = emeasure ?D X"
using X `J \<subseteq> K` apply (subst emeasure_distr)
by (auto intro!: measurable_restrict_subset simp: space_PiM)
qed
lemma (in product_prob_space) emeasure_prod_emb[simp]:
- assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)"
- shows "emeasure (Pi\<^isub>M L M) (prod_emb L M J X) = emeasure (Pi\<^isub>M J M) X"
+ assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^sub>M J M)"
+ shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X"
by (subst distr_restrict[OF L])
(simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
definition
limP :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i set \<Rightarrow> ('i \<Rightarrow> 'a) measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
- "limP I M P = extend_measure (\<Pi>\<^isub>E i\<in>I. space (M i))
+ "limP I M P = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
{(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
- (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j))
- (\<lambda>(J, X). emeasure (P J) (Pi\<^isub>E J X))"
+ (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
+ (\<lambda>(J, X). emeasure (P J) (Pi\<^sub>E J X))"
-abbreviation "lim\<^isub>P \<equiv> limP"
+abbreviation "lim\<^sub>P \<equiv> limP"
lemma space_limP[simp]: "space (limP I M P) = space (PiM I M)"
by (auto simp add: limP_def space_PiM prod_emb_def intro!: space_extend_measure)
@@ -71,10 +71,10 @@
lemma sets_limP[simp]: "sets (limP I M P) = sets (PiM I M)"
by (auto simp add: limP_def sets_PiM prod_algebra_def prod_emb_def intro!: sets_extend_measure)
-lemma measurable_limP1[simp]: "measurable (limP I M P) M' = measurable (\<Pi>\<^isub>M i\<in>I. M i) M'"
+lemma measurable_limP1[simp]: "measurable (limP I M P) M' = measurable (\<Pi>\<^sub>M i\<in>I. M i) M'"
unfolding measurable_def by auto
-lemma measurable_limP2[simp]: "measurable M' (limP I M P) = measurable M' (\<Pi>\<^isub>M i\<in>I. M i)"
+lemma measurable_limP2[simp]: "measurable M' (limP I M P) = measurable M' (\<Pi>\<^sub>M i\<in>I. M i)"
unfolding measurable_def by auto
locale projective_family =
@@ -90,14 +90,14 @@
assumes "finite J"
assumes "J \<subseteq> I"
assumes A: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> sets (M i)"
- shows "emeasure (limP J M P) (Pi\<^isub>E J A) = emeasure (P J) (Pi\<^isub>E J A)"
+ shows "emeasure (limP J M P) (Pi\<^sub>E J A) = emeasure (P J) (Pi\<^sub>E J A)"
proof -
- have "Pi\<^isub>E J (restrict A J) \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
+ have "Pi\<^sub>E J (restrict A J) \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))"
using sets.sets_into_space[OF A] by (auto simp: PiE_iff) blast
- hence "emeasure (limP J M P) (Pi\<^isub>E J A) =
- emeasure (limP J M P) (prod_emb J M J (Pi\<^isub>E J A))"
+ hence "emeasure (limP J M P) (Pi\<^sub>E J A) =
+ emeasure (limP J M P) (prod_emb J M J (Pi\<^sub>E J A))"
using assms(1-3) sets.sets_into_space by (auto simp add: prod_emb_id PiE_def Pi_def)
- also have "\<dots> = emeasure (P J) (Pi\<^isub>E J A)"
+ also have "\<dots> = emeasure (P J) (Pi\<^sub>E J A)"
proof (rule emeasure_extend_measure_Pair[OF limP_def])
show "positive (sets (limP J M P)) (P J)" unfolding positive_def by auto
show "countably_additive (sets (limP J M P)) (P J)" unfolding countably_additive_def
@@ -105,10 +105,10 @@
show "(J \<noteq> {} \<or> J = {}) \<and> finite J \<and> J \<subseteq> J \<and> A \<in> (\<Pi> j\<in>J. sets (M j))"
using assms by auto
fix K and X::"'i \<Rightarrow> 'a set"
- show "prod_emb J M K (Pi\<^isub>E K X) \<in> Pow (\<Pi>\<^isub>E i\<in>J. space (M i))"
+ show "prod_emb J M K (Pi\<^sub>E K X) \<in> Pow (\<Pi>\<^sub>E i\<in>J. space (M i))"
by (auto simp: prod_emb_def)
assume JX: "(K \<noteq> {} \<or> J = {}) \<and> finite K \<and> K \<subseteq> J \<and> X \<in> (\<Pi> j\<in>K. sets (M j))"
- thus "emeasure (P J) (prod_emb J M K (Pi\<^isub>E K X)) = emeasure (P K) (Pi\<^isub>E K X)"
+ thus "emeasure (P J) (prod_emb J M K (Pi\<^sub>E K X)) = emeasure (P K) (Pi\<^sub>E K X)"
using assms
apply (cases "J = {}")
apply (simp add: prod_emb_id)
@@ -123,16 +123,16 @@
assumes "J \<subseteq> I"
shows "limP J M P = P J" (is "?P = _")
proof (rule measure_eqI_generator_eq)
- let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
- let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
+ let ?J = "{Pi\<^sub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
+ let ?\<Omega> = "(\<Pi>\<^sub>E k\<in>J. space (M k))"
interpret prob_space "P J" using proj_prob_space `finite J` by simp
- show "emeasure ?P (\<Pi>\<^isub>E k\<in>J. space (M k)) \<noteq> \<infinity>" using assms `finite J` by (auto simp: emeasure_limP)
+ show "emeasure ?P (\<Pi>\<^sub>E k\<in>J. space (M k)) \<noteq> \<infinity>" using assms `finite J` by (auto simp: emeasure_limP)
show "sets (limP J M P) = sigma_sets ?\<Omega> ?J" "sets (P J) = sigma_sets ?\<Omega> ?J"
using `finite J` proj_sets by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
fix X assume "X \<in> ?J"
- then obtain E where X: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
+ then obtain E where X: "X = Pi\<^sub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
with `finite J` have "X \<in> sets (limP J M P)" by simp
- have emb_self: "prod_emb J M J (Pi\<^isub>E J E) = Pi\<^isub>E J E"
+ have emb_self: "prod_emb J M J (Pi\<^sub>E J E) = Pi\<^sub>E J E"
using E sets.sets_into_space
by (auto intro!: prod_emb_PiE_same_index)
show "emeasure (limP J M P) X = emeasure (P J) X"
@@ -150,11 +150,11 @@
"emb L K X \<equiv> prod_emb L M K X"
lemma prod_emb_injective:
- assumes "J \<subseteq> L" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
+ assumes "J \<subseteq> L" and sets: "X \<in> sets (Pi\<^sub>M J M)" "Y \<in> sets (Pi\<^sub>M J M)"
assumes "emb L J X = emb L J Y"
shows "X = Y"
proof (rule injective_vimage_restrict)
- show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
+ show "X \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))"
using sets[THEN sets.sets_into_space] by (auto simp: space_PiM)
have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
proof
@@ -163,20 +163,20 @@
from not_empty show "\<exists>x. x \<in> space (M i)" by (auto simp add: proj_space space_PiM)
qed
from bchoice[OF this]
- show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by (auto simp: PiE_def)
- show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
+ show "(\<Pi>\<^sub>E i\<in>L. space (M i)) \<noteq> {}" by (auto simp: PiE_def)
+ show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i))"
using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)
qed fact
definition generator :: "('i \<Rightarrow> 'a) set set" where
- "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))"
+ "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^sub>M J M))"
lemma generatorI':
- "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> generator"
+ "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emb I J X \<in> generator"
unfolding generator_def by auto
lemma algebra_generator:
- assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
+ assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
unfolding algebra_def algebra_axioms_def ring_of_sets_iff
proof (intro conjI ballI)
let ?G = generator
@@ -187,13 +187,13 @@
by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
simp: sigma_sets.Empty generator_def prod_emb_def)
from `i \<in> I` show "?\<Omega> \<in> ?G"
- by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
+ by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^sub>E {i} (\<lambda>i. space (M i))"]
simp: generator_def prod_emb_def)
fix A assume "A \<in> ?G"
- then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA"
+ then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^sub>M JA M)" and A: "A = emb I JA XA"
by (auto simp: generator_def)
fix B assume "B \<in> ?G"
- then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB"
+ then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^sub>M JB M)" and B: "B = emb I JB XB"
by (auto simp: generator_def)
let ?RA = "emb (JA \<union> JB) JA XA"
let ?RB = "emb (JA \<union> JB) JB XB"
@@ -204,7 +204,7 @@
qed
lemma sets_PiM_generator:
- "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+ "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) generator"
proof cases
assume "I = {}" then show ?thesis
unfolding generator_def
@@ -213,7 +213,7 @@
assume "I \<noteq> {}"
show ?thesis
proof
- show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+ show "sets (Pi\<^sub>M I M) \<subseteq> sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) generator"
unfolding sets_PiM
proof (safe intro!: sigma_sets_subseteq)
fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
@@ -223,19 +223,19 @@
qed
lemma generatorI:
- "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
+ "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
unfolding generator_def by auto
definition mu_G ("\<mu>G") where
"\<mu>G A =
- (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (limP J M P) X))"
+ (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^sub>M J M). A = emb I J X \<longrightarrow> x = emeasure (limP J M P) X))"
lemma mu_G_spec:
- assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
+ assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^sub>M J M)"
shows "\<mu>G A = emeasure (limP J M P) X"
unfolding mu_G_def
proof (intro the_equality allI impI ballI)
- fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)"
+ fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^sub>M K M)"
have "emeasure (limP K M P) Y = emeasure (limP (K \<union> J) M P) (emb (K \<union> J) K Y)"
using K J by simp
also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
@@ -246,31 +246,31 @@
qed (insert J, force)
lemma mu_G_eq:
- "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (limP J M P) X"
+ "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (limP J M P) X"
by (intro mu_G_spec) auto
lemma generator_Ex:
assumes *: "A \<in> generator"
- shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (limP J M P) X"
+ shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (limP J M P) X"
proof -
- from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
+ from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^sub>M J M)"
unfolding generator_def by auto
with mu_G_spec[OF this] show ?thesis by auto
qed
lemma generatorE:
assumes A: "A \<in> generator"
- obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (limP J M P) X"
+ obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (limP J M P) X"
using generator_Ex[OF A] by atomize_elim auto
lemma merge_sets:
- "J \<inter> K = {} \<Longrightarrow> A \<in> sets (Pi\<^isub>M (J \<union> K) M) \<Longrightarrow> x \<in> space (Pi\<^isub>M J M) \<Longrightarrow> (\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
+ "J \<inter> K = {} \<Longrightarrow> A \<in> sets (Pi\<^sub>M (J \<union> K) M) \<Longrightarrow> x \<in> space (Pi\<^sub>M J M) \<Longrightarrow> (\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^sub>M K M) \<in> sets (Pi\<^sub>M K M)"
by simp
lemma merge_emb:
- assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"
- shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
- emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
+ assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^sub>M J M)"
+ shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) =
+ emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M))"
proof -
have [simp]: "\<And>x J K L. merge J K (y, restrict x L) = merge J (K \<inter> L) (y, x)"
by (auto simp: restrict_def merge_def)
@@ -288,7 +288,7 @@
assumes "I \<noteq> {}"
shows "positive generator \<mu>G"
proof -
- interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
+ interpret G!: algebra "\<Pi>\<^sub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
show ?thesis
proof (intro positive_def[THEN iffD2] conjI ballI)
from generatorE[OF G.empty_sets] guess J X . note this[simp]
@@ -306,7 +306,7 @@
assumes "I \<noteq> {}"
shows "additive generator \<mu>G"
proof -
- interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
+ interpret G!: algebra "\<Pi>\<^sub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
show ?thesis
proof (intro additive_def[THEN iffD2] ballI impI)
fix A assume "A \<in> generator" with generatorE guess J X . note J = this
@@ -337,12 +337,12 @@
proof
fix J::"'i set" assume "finite J"
interpret f: finite_product_prob_space M J proof qed fact
- show "emeasure (Pi\<^isub>M J M) (space (Pi\<^isub>M J M)) \<noteq> \<infinity>" by simp
- show "\<exists>A. range A \<subseteq> sets (Pi\<^isub>M J M) \<and>
- (\<Union>i. A i) = space (Pi\<^isub>M J M) \<and>
- (\<forall>i. emeasure (Pi\<^isub>M J M) (A i) \<noteq> \<infinity>)" using sigma_finite[OF `finite J`]
+ show "emeasure (Pi\<^sub>M J M) (space (Pi\<^sub>M J M)) \<noteq> \<infinity>" by simp
+ show "\<exists>A. range A \<subseteq> sets (Pi\<^sub>M J M) \<and>
+ (\<Union>i. A i) = space (Pi\<^sub>M J M) \<and>
+ (\<forall>i. emeasure (Pi\<^sub>M J M) (A i) \<noteq> \<infinity>)" using sigma_finite[OF `finite J`]
by (auto simp add: sigma_finite_measure_def)
- show "emeasure (Pi\<^isub>M J M) (space (Pi\<^isub>M J M)) = 1" by (rule f.emeasure_space_1)
+ show "emeasure (Pi\<^sub>M J M) (space (Pi\<^sub>M J M)) = 1" by (rule f.emeasure_space_1)
qed simp_all
lemma (in product_prob_space) limP_PiM_finite[simp]:
--- a/src/HOL/Probability/Projective_Limit.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy Tue Aug 13 16:25:47 2013 +0200
@@ -17,15 +17,15 @@
subsection {* Sequences of Finite Maps in Compact Sets *}
locale finmap_seqs_into_compact =
- fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^isub>F 'a)" and M
+ fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a)" and M
assumes compact: "\<And>n. compact (K n)"
assumes f_in_K: "\<And>n. K n \<noteq> {}"
assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)"
assumes proj_in_K:
- "\<And>t n m. m \<ge> n \<Longrightarrow> t \<in> domain (f n) \<Longrightarrow> (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n"
+ "\<And>t n m. m \<ge> n \<Longrightarrow> t \<in> domain (f n) \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"
begin
-lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n)"
+lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n)"
using proj_in_K f_in_K
proof cases
obtain k where "k \<in> K (Suc 0)" using f_in_K by auto
@@ -36,11 +36,11 @@
qed blast
lemma proj_in_KE:
- obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K n"
+ obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"
using proj_in_K' by blast
lemma compact_projset:
- shows "compact ((\<lambda>k. (k)\<^isub>F i) ` K n)"
+ shows "compact ((\<lambda>k. (k)\<^sub>F i) ` K n)"
using continuous_proj compact by (rule compact_continuous_image)
end
@@ -58,38 +58,38 @@
thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) ----> l" by blast
qed
-sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^isub>F n) ----> l)"
+sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^sub>F n) ----> l)"
proof
fix n s
assume "subseq s"
from proj_in_KE[of n] guess n0 . note n0 = this
- have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0"
+ have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0"
proof safe
fix i assume "n0 \<le> i"
also have "\<dots> \<le> s i" by (rule seq_suble) fact
finally have "n0 \<le> s i" .
- with n0 show "((f \<circ> s) i)\<^isub>F n \<in> (\<lambda>k. (k)\<^isub>F n) ` K n0 "
+ with n0 show "((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0 "
by auto
qed
from compactE'[OF compact_projset this] guess ls rs .
- thus "\<exists>r'. subseq r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^isub>F n) ----> l)" by (auto simp: o_def)
+ thus "\<exists>r'. subseq r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^sub>F n) ----> l)" by (auto simp: o_def)
qed
-lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^isub>F n) ----> l"
+lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^sub>F n) ----> l"
proof -
- obtain l where "(\<lambda>i. ((f o (diagseq o op + (Suc n))) i)\<^isub>F n) ----> l"
+ obtain l where "(\<lambda>i. ((f o (diagseq o op + (Suc n))) i)\<^sub>F n) ----> l"
proof (atomize_elim, rule diagseq_holds)
fix r s n
assume "subseq r"
- assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^isub>F n) ----> l"
- then obtain l where "((\<lambda>i. (f i)\<^isub>F n) o s) ----> l"
+ assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) ----> l"
+ then obtain l where "((\<lambda>i. (f i)\<^sub>F n) o s) ----> l"
by (auto simp: o_def)
- hence "((\<lambda>i. (f i)\<^isub>F n) o s o r) ----> l" using `subseq r`
+ hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) ----> l" using `subseq r`
by (rule LIMSEQ_subseq_LIMSEQ)
- thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^isub>F n) ----> l" by (auto simp add: o_def)
+ thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^sub>F n) ----> l" by (auto simp add: o_def)
qed
- hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^isub>F n) ----> l" by (simp add: ac_simps)
- hence "(\<lambda>i. (f (diagseq i))\<^isub>F n) ----> l" by (rule LIMSEQ_offset)
+ hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^sub>F n) ----> l" by (simp add: ac_simps)
+ hence "(\<lambda>i. (f (diagseq i))\<^sub>F n) ----> l" by (rule LIMSEQ_offset)
thus ?thesis ..
qed
@@ -101,14 +101,14 @@
for I::"'i set" and P
begin
-abbreviation "lim\<^isub>B \<equiv> (\<lambda>J P. limP J (\<lambda>_. borel) P)"
+abbreviation "lim\<^sub>B \<equiv> (\<lambda>J P. limP J (\<lambda>_. borel) P)"
lemma emeasure_limB_emb_not_empty:
assumes "I \<noteq> {}"
assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
- shows "emeasure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = emeasure (lim\<^isub>B J P) (Pi\<^isub>E J B)"
+ shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)"
proof -
- let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space borel"
+ let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space borel"
let ?G = generator
interpret G!: algebra ?\<Omega> generator by (intro algebra_generator) fact
note mu_G_mono =
@@ -135,17 +135,17 @@
using Z positive_mu_G[OF `I \<noteq> {}`] by (auto intro!: INF_greatest simp: positive_def)
ultimately have "0 < ?a" by auto
hence "?a \<noteq> -\<infinity>" by auto
- have "\<forall>n. \<exists>J B. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> B \<in> sets (Pi\<^isub>M J (\<lambda>_. borel)) \<and>
- Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (lim\<^isub>B J P) B"
+ have "\<forall>n. \<exists>J B. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> B \<in> sets (Pi\<^sub>M J (\<lambda>_. borel)) \<and>
+ Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (lim\<^sub>B J P) B"
using Z by (intro allI generator_Ex) auto
then obtain J' B' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I"
- "\<And>n. B' n \<in> sets (\<Pi>\<^isub>M i\<in>J' n. borel)"
+ "\<And>n. B' n \<in> sets (\<Pi>\<^sub>M i\<in>J' n. borel)"
and Z_emb: "\<And>n. Z n = emb I (J' n) (B' n)"
unfolding choice_iff by blast
moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
moreover def B \<equiv> "\<lambda>n. emb (J n) (J' n) (B' n)"
ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I"
- "\<And>n. B n \<in> sets (\<Pi>\<^isub>M i\<in>J n. borel)"
+ "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)"
by auto
have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
unfolding J_def by force
@@ -230,8 +230,8 @@
"\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
"\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
unfolding choice_iff by blast
- def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
- have K_sets: "\<And>n. K n \<in> sets (Pi\<^isub>M (J n) (\<lambda>_. borel))"
+ def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
+ have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
unfolding K_def
using compact_imp_closed[OF `compact (K' _)`]
by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
@@ -250,7 +250,7 @@
unfolding Z_eq unfolding Z'_def
proof (rule prod_emb_mono, safe)
fix n x assume "x \<in> K n"
- hence "fm n x \<in> K' n" "x \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))"
+ hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
by (simp_all add: K_def proj_space)
note this(1)
also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
@@ -259,7 +259,7 @@
proof safe
fix y assume "y \<in> B n"
moreover
- hence "y \<in> space (Pi\<^isub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]
+ hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]
by (auto simp add: proj_space proj_sets)
assume "fm n x = fm n y"
note inj_onD[OF inj_on_fm[OF space_borel],
@@ -297,7 +297,7 @@
by (subst mu_G_eq) (auto simp: limP_finite proj_sets mu_G_eq)
interpret finite_measure "(limP (J n) (\<lambda>_. borel) P)"
proof
- have "emeasure (limP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^isub>E space borel) \<noteq> \<infinity>"
+ have "emeasure (limP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^sub>E space borel) \<noteq> \<infinity>"
using J by (subst emeasure_limP) auto
thus "emeasure (limP (J n) (\<lambda>_. borel) P) (space (limP (J n) (\<lambda>_. borel) P)) \<noteq> \<infinity>"
by (simp add: space_PiM)
@@ -414,15 +414,15 @@
hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
by (intro fm_in_K') simp_all
- show "(fm (Suc m) (y (Suc m)))\<^isub>F t \<in> (\<lambda>k. (k)\<^isub>F t) ` K' (Suc n)"
+ show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"
apply (rule image_eqI[OF _ img])
using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
unfolding j by (subst proj_fm, auto)+
qed
- have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z"
+ have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z"
using diagonal_tendsto ..
then obtain z where z:
- "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
+ "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
unfolding choice_iff by blast
{
fix n :: nat assume "n \<ge> 1"
@@ -434,14 +434,14 @@
assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
hence "t \<in> Utn ` J n" by simp
then obtain j where j: "t = Utn j" "j \<in> J n" by auto
- have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> z t"
+ have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
apply (subst (2) tendsto_iff, subst eventually_sequentially)
proof safe
fix e :: real assume "0 < e"
{ fix i x assume "i \<ge> n" "t \<in> domain (fm n x)"
moreover
hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
- ultimately have "(fm i x)\<^isub>F t = (fm n x)\<^isub>F t"
+ ultimately have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t"
using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn])
} note index_shift = this
have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
@@ -450,21 +450,21 @@
apply (rule seq_suble[OF subseq_diagseq])
done
from z
- have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e"
+ have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e"
unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
- dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^isub>F t) (z t) < e" by auto
- show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e "
+ dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto
+ show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e "
proof (rule exI[where x="max N n"], safe)
fix na assume "max N n \<le> na"
- hence "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) =
- dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^isub>F t) (z t)" using t
+ hence "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) =
+ dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t
by (subst index_shift[OF I]) auto
also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp
- finally show "dist ((fm n (y (Suc (diagseq na))))\<^isub>F t) (z t) < e" .
+ finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" .
qed
qed
- hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^isub>F t) ----> (finmap_of (Utn ` J n) z)\<^isub>F t"
+ hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> (finmap_of (Utn ` J n) z)\<^sub>F t"
by (simp add: tendsto_intros)
} ultimately
have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
@@ -491,7 +491,7 @@
moreover hence "from_nat_into (\<Union>n. J n) (Utn i) = i"
unfolding Utn_def
by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto
- ultimately show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^isub>F (Utn i)"
+ ultimately show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)"
by (simp add: finmap_eq_iff fm_def compose_def)
qed
finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
@@ -514,27 +514,27 @@
qed
then guess \<mu> .. note \<mu> = this
def f \<equiv> "finmap_of J B"
- show "emeasure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = emeasure (lim\<^isub>B J P) (Pi\<^isub>E J B)"
+ show "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)"
proof (subst emeasure_extend_measure_Pair[OF limP_def, of I "\<lambda>_. borel" \<mu>])
- show "positive (sets (lim\<^isub>B I P)) \<mu>" "countably_additive (sets (lim\<^isub>B I P)) \<mu>"
+ show "positive (sets (lim\<^sub>B I P)) \<mu>" "countably_additive (sets (lim\<^sub>B I P)) \<mu>"
using \<mu> unfolding sets_limP sets_PiM_generator by (auto simp: measure_space_def)
next
show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> B \<in> J \<rightarrow> sets borel"
using assms by (auto simp: f_def)
next
fix J and X::"'i \<Rightarrow> 'a set"
- show "prod_emb I (\<lambda>_. borel) J (Pi\<^isub>E J X) \<in> Pow (I \<rightarrow>\<^isub>E space borel)"
+ show "prod_emb I (\<lambda>_. borel) J (Pi\<^sub>E J X) \<in> Pow (I \<rightarrow>\<^sub>E space borel)"
by (auto simp: prod_emb_def)
assume JX: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> J \<rightarrow> sets borel"
- hence "emb I J (Pi\<^isub>E J X) \<in> generator" using assms
- by (intro generatorI[where J=J and X="Pi\<^isub>E J X"]) (auto intro: sets_PiM_I_finite)
- hence "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))" using \<mu> by simp
- also have "\<dots> = emeasure (P J) (Pi\<^isub>E J X)"
+ hence "emb I J (Pi\<^sub>E J X) \<in> generator" using assms
+ by (intro generatorI[where J=J and X="Pi\<^sub>E J X"]) (auto intro: sets_PiM_I_finite)
+ hence "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))" using \<mu> by simp
+ also have "\<dots> = emeasure (P J) (Pi\<^sub>E J X)"
using JX assms proj_sets
by (subst mu_G_eq) (auto simp: mu_G_eq limP_finite intro: sets_PiM_I_finite)
- finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = emeasure (P J) (Pi\<^isub>E J X)" .
+ finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = emeasure (P J) (Pi\<^sub>E J X)" .
next
- show "emeasure (P J) (Pi\<^isub>E J B) = emeasure (limP J (\<lambda>_. borel) P) (Pi\<^isub>E J B)"
+ show "emeasure (P J) (Pi\<^sub>E J B) = emeasure (limP J (\<lambda>_. borel) P) (Pi\<^sub>E J B)"
using assms by (simp add: f_def limP_finite Pi_def)
qed
qed
@@ -542,7 +542,7 @@
end
hide_const (open) PiF
-hide_const (open) Pi\<^isub>F
+hide_const (open) Pi\<^sub>F
hide_const (open) Pi'
hide_const (open) Abs_finmap
hide_const (open) Rep_finmap
@@ -551,9 +551,9 @@
hide_const (open) domain
hide_const (open) basis_finmap
-sublocale polish_projective \<subseteq> P!: prob_space "(lim\<^isub>B I P)"
+sublocale polish_projective \<subseteq> P!: prob_space "(lim\<^sub>B I P)"
proof
- show "emeasure (lim\<^isub>B I P) (space (lim\<^isub>B I P)) = 1"
+ show "emeasure (lim\<^sub>B I P) (space (lim\<^sub>B I P)) = 1"
proof cases
assume "I = {}"
interpret prob_space "P {}" using proj_prob_space by simp
@@ -563,7 +563,7 @@
assume "I \<noteq> {}"
then obtain i where "i \<in> I" by auto
interpret prob_space "P {i}" using proj_prob_space by simp
- have R: "(space (lim\<^isub>B I P)) = (emb I {i} (Pi\<^isub>E {i} (\<lambda>_. space borel)))"
+ have R: "(space (lim\<^sub>B I P)) = (emb I {i} (Pi\<^sub>E {i} (\<lambda>_. space borel)))"
by (auto simp: prod_emb_def space_PiM)
moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM PiE_def)
ultimately show ?thesis using `i \<in> I`
@@ -578,13 +578,13 @@
lemma emeasure_limB_emb:
assumes X: "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
- shows "emeasure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = emeasure (P J) (Pi\<^isub>E J B)"
+ shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (P J) (Pi\<^sub>E J B)"
proof cases
interpret prob_space "P {}" using proj_prob_space by simp
assume "J = {}"
- moreover have "emb I {} {\<lambda>x. undefined} = space (lim\<^isub>B I P)"
+ moreover have "emb I {} {\<lambda>x. undefined} = space (lim\<^sub>B I P)"
by (auto simp: space_PiM prod_emb_def)
- moreover have "{\<lambda>x. undefined} = space (lim\<^isub>B {} P)"
+ moreover have "{\<lambda>x. undefined} = space (lim\<^sub>B {} P)"
by (auto simp: space_PiM prod_emb_def)
ultimately show ?thesis
by (simp add: P.emeasure_space_1 limP_finite emeasure_space_1 del: space_limP)
@@ -595,7 +595,7 @@
lemma measure_limB_emb:
assumes "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
- shows "measure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = measure (P J) (Pi\<^isub>E J B)"
+ shows "measure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = measure (P J) (Pi\<^sub>E J B)"
proof -
interpret prob_space "P J" using proj_prob_space assms by simp
show ?thesis
@@ -613,7 +613,7 @@
proof qed
lemma (in polish_product_prob_space) limP_eq_PiM:
- "I \<noteq> {} \<Longrightarrow> lim\<^isub>P I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
+ "I \<noteq> {} \<Longrightarrow> lim\<^sub>P I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
PiM I (\<lambda>_. borel)"
by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_limB_emb)
--- a/src/HOL/Probability/Radon_Nikodym.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy Tue Aug 13 16:25:47 2013 +0200
@@ -44,7 +44,7 @@
qed fact
lemma (in sigma_finite_measure) Ex_finite_integrable_function:
- shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
+ shows "\<exists>h\<in>borel_measurable M. integral\<^sup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
proof -
obtain A :: "nat \<Rightarrow> 'a set" where
range[measurable]: "range A \<subseteq> sets M" and
@@ -67,7 +67,7 @@
proof (safe intro!: bexI[of _ ?h] del: notI)
have "\<And>i. A i \<in> sets M"
using range by fastforce+
- then have "integral\<^isup>P M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
+ then have "integral\<^sup>P M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
by (simp add: positive_integral_suminf positive_integral_cmult_indicator)
also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)"
proof (rule suminf_le_pos)
@@ -82,7 +82,7 @@
finally show "n N * emeasure M (A N) \<le> (1 / 2) ^ Suc N" .
show "0 \<le> n N * emeasure M (A N)" using n[of N] `A N \<in> sets M` by (simp add: emeasure_nonneg)
qed
- finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
+ finally show "integral\<^sup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
next
{ fix x assume "x \<in> space M"
then obtain i where "x \<in> A i" using space[symmetric] by auto
@@ -294,7 +294,7 @@
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
proof -
interpret N: finite_measure N by fact
- def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
+ def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
{ fix f have "f \<in> G \<Longrightarrow> f \<in> borel_measurable M" by (auto simp: G_def) }
note this[measurable_dest]
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
@@ -313,16 +313,16 @@
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
by (auto simp: indicator_def max_def)
- hence "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) =
- (\<integral>\<^isup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +
- (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
+ hence "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) =
+ (\<integral>\<^sup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +
+ (\<integral>\<^sup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
using f g sets unfolding G_def
by (auto cong: positive_integral_cong intro!: positive_integral_add)
also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"
using f g sets unfolding G_def by (auto intro!: add_mono)
also have "\<dots> = N A"
using plus_emeasure[OF sets'] union by auto
- finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
+ finally show "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
next
fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max)
qed }
@@ -336,33 +336,33 @@
using f by (auto simp: G_def intro: SUP_upper2) }
next
fix A assume "A \<in> sets M"
- have "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
- (\<integral>\<^isup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
+ have "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
+ (\<integral>\<^sup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
by (intro positive_integral_cong) (simp split: split_indicator)
- also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))"
+ also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i x * indicator A x \<partial>M))"
using `incseq f` f `A \<in> sets M`
by (intro positive_integral_monotone_convergence_SUP)
(auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
- finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
+ finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def)
qed }
note SUP_in_G = this
- let ?y = "SUP g : G. integral\<^isup>P M g"
+ let ?y = "SUP g : G. integral\<^sup>P M g"
have y_le: "?y \<le> N (space M)" unfolding G_def
proof (safe intro!: SUP_least)
- fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A"
- from this[THEN bspec, OF sets.top] show "integral\<^isup>P M g \<le> N (space M)"
+ fix g assume "\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A"
+ from this[THEN bspec, OF sets.top] show "integral\<^sup>P M g \<le> N (space M)"
by (simp cong: positive_integral_cong)
qed
- from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^isup>P M"] guess ys .. note ys = this
- then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n"
+ from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^sup>P M"] guess ys .. note ys = this
+ then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>P M g = ys n"
proof safe
- fix n assume "range ys \<subseteq> integral\<^isup>P M ` G"
- hence "ys n \<in> integral\<^isup>P M ` G" by auto
- thus "\<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" by auto
+ fix n assume "range ys \<subseteq> integral\<^sup>P M ` G"
+ hence "ys n \<in> integral\<^sup>P M ` G" by auto
+ thus "\<exists>g. g\<in>G \<and> integral\<^sup>P M g = ys n" by auto
qed
- from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^isup>P M (gs n) = ys n" by auto
- hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto
+ from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^sup>P M (gs n) = ys n" by auto
+ hence y_eq: "?y = (SUP i. integral\<^sup>P M (gs i))" using ys by auto
let ?g = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})"
def f \<equiv> "\<lambda>x. SUP i. ?g i x"
let ?F = "\<lambda>A x. f x * indicator A x"
@@ -380,23 +380,23 @@
by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)
from SUP_in_G[OF this g_in_G] have [measurable]: "f \<in> G" unfolding f_def .
then have [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
- have "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding f_def
+ have "integral\<^sup>P M f = (SUP i. integral\<^sup>P M (?g i))" unfolding f_def
using g_in_G `incseq ?g`
by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
also have "\<dots> = ?y"
proof (rule antisym)
- show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
+ show "(SUP i. integral\<^sup>P M (?g i)) \<le> ?y"
using g_in_G by (auto intro: Sup_mono simp: SUP_def)
- show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
+ show "?y \<le> (SUP i. integral\<^sup>P M (?g i))" unfolding y_eq
by (auto intro!: SUP_mono positive_integral_mono Max_ge)
qed
- finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
+ finally have int_f_eq_y: "integral\<^sup>P M f = ?y" .
have "\<And>x. 0 \<le> f x"
unfolding f_def using `\<And>i. gs i \<in> G`
by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def)
- let ?t = "\<lambda>A. N A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
+ let ?t = "\<lambda>A. N A - (\<integral>\<^sup>+x. ?F A x \<partial>M)"
let ?M = "diff_measure N (density M f)"
- have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A"
+ have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A"
using `f \<in> G` unfolding G_def by auto
have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure ?M A = ?t A"
proof (subst emeasure_diff_measure)
@@ -415,8 +415,8 @@
fix A assume A: "A \<in> null_sets M"
with `absolutely_continuous M N` have "A \<in> null_sets N"
unfolding absolutely_continuous_def by auto
- moreover with A have "(\<integral>\<^isup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)
- ultimately have "N A - (\<integral>\<^isup>+ x. ?F A x \<partial>M) = 0"
+ moreover with A have "(\<integral>\<^sup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)
+ ultimately have "N A - (\<integral>\<^sup>+ x. ?F A x \<partial>M) = 0"
using positive_integral_positive[of M] by (auto intro!: antisym)
then show "A \<in> null_sets ?M"
using A by (simp add: emeasure_M null_sets_def sets_eq)
@@ -436,9 +436,9 @@
then have pos_M: "0 < emeasure M (space M)"
using emeasure_nonneg[of M "space M"] by (simp add: le_less)
moreover
- have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
+ have "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
using `f \<in> G` unfolding G_def by auto
- hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
+ hence "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
using M'.finite_emeasure_space by auto
moreover
def b \<equiv> "?M (space M) / emeasure M (space M) / 2"
@@ -460,32 +460,32 @@
let ?f0 = "\<lambda>x. f x + b * indicator A0 x"
{ fix A assume A: "A \<in> sets M"
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
- have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
- (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
+ have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =
+ (\<integral>\<^sup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
by (auto intro!: positive_integral_cong split: split_indicator)
- hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
- (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"
+ hence "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =
+ (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"
using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`
by (simp add: positive_integral_add positive_integral_cmult_indicator G_def) }
note f0_eq = this
{ fix A assume A: "A \<in> sets M"
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
- have f_le_v: "(\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto
+ have f_le_v: "(\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto
note f0_eq[OF A]
- also have "(\<integral>\<^isup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^isup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"
+ also have "(\<integral>\<^sup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^sup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"
using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
by (auto intro!: add_left_mono)
- also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?M A"
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + ?M A"
using emeasure_mono[of "A \<inter> A0" A ?M] `A \<in> sets M` `A0 \<in> sets M`
by (auto intro!: add_left_mono simp: sets_eq)
also have "\<dots> \<le> N A"
unfolding emeasure_M[OF `A \<in> sets M`]
using f_le_v N.emeasure_eq_measure[of A] positive_integral_positive[of M "?F A"]
- by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "N A") auto
- finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
+ by (cases "\<integral>\<^sup>+x. ?F A x \<partial>M", cases "N A") auto
+ finally have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G`
by (auto intro!: ereal_add_nonneg_nonneg simp: G_def)
- have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
+ have int_f_finite: "integral\<^sup>P M f \<noteq> \<infinity>"
by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le)
have "0 < ?M (space M) - emeasure ?Mb (space M)"
using pos_t
@@ -504,13 +504,13 @@
by (auto simp: absolutely_continuous_def null_sets_def)
then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto
hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff)
- with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * emeasure M A0" unfolding int_f_eq_y
+ with int_f_finite have "?y + 0 < integral\<^sup>P M f + b * emeasure M A0" unfolding int_f_eq_y
using `f \<in> G`
by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive)
- also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF sets.top] `A0 \<in> sets M` sets.sets_into_space
+ also have "\<dots> = integral\<^sup>P M ?f0" using f0_eq[OF sets.top] `A0 \<in> sets M` sets.sets_into_space
by (simp cong: positive_integral_cong)
- finally have "?y < integral\<^isup>P M ?f0" by simp
- moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
+ finally have "?y < integral\<^sup>P M ?f0" by simp
+ moreover from `?f0 \<in> G` have "integral\<^sup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
ultimately show False by auto
qed
let ?f = "\<lambda>x. max 0 (f x)"
@@ -521,7 +521,7 @@
fix A assume A: "A\<in>sets (density M ?f)"
then show "emeasure (density M ?f) A = emeasure N A"
using `f \<in> G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]
- by (cases "integral\<^isup>P M (?F A)")
+ by (cases "integral\<^sup>P M (?F A)")
(auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric])
qed auto
qed
@@ -686,12 +686,12 @@
and f_density: "\<And>i. density (?M i) (f i) = ?N i"
by auto
{ fix A i assume A: "A \<in> sets M"
- with Q borel have "(\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
+ with Q borel have "(\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
by (auto simp add: emeasure_density positive_integral_density subset_eq
intro!: positive_integral_cong split: split_indicator)
also have "\<dots> = emeasure N (Q i \<inter> A)"
using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)
- finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
+ finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
note integral_eq = this
let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"
show ?thesis
@@ -707,15 +707,15 @@
have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
"\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
- have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
+ have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
using borel by (intro positive_integral_cong) (auto simp: indicator_def)
- also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
+ also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
using borel Qi Q0(1) `A \<in> sets M`
by (subst positive_integral_add) (auto simp del: ereal_infty_mult
simp add: positive_integral_cmult_indicator sets.Int intro!: suminf_0_le)
also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"
by (subst integral_eq[OF `A \<in> sets M`], subst positive_integral_suminf) auto
- finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
+ finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
using Q Q_sets `A \<in> sets M`
by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
@@ -728,7 +728,7 @@
using Q_sets `A \<in> sets M` Q0(1) by auto
moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
using `A \<in> sets M` sets.sets_into_space Q0 by auto
- ultimately have "N A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
+ ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)"
using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "Q0 \<inter> A"] by (simp add: sets_eq)
with `A \<in> sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A"
by (auto simp: subset_eq emeasure_density)
@@ -741,12 +741,12 @@
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
proof -
from Ex_finite_integrable_function
- obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and
+ obtain h where finite: "integral\<^sup>P M h \<noteq> \<infinity>" and
borel: "h \<in> borel_measurable M" and
nn: "\<And>x. 0 \<le> h x" and
pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
"\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
- let ?T = "\<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
+ let ?T = "\<lambda>A. (\<integral>\<^sup>+x. h x * indicator A x \<partial>M)"
let ?MT = "density M h"
from borel finite nn interpret T: finite_measure ?MT
by (auto intro!: finite_measureI cong: positive_integral_cong simp: emeasure_density)
@@ -773,28 +773,28 @@
lemma finite_density_unique:
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
- and fin: "integral\<^isup>P M f \<noteq> \<infinity>"
+ and fin: "integral\<^sup>P M f \<noteq> \<infinity>"
shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
proof (intro iffI ballI)
fix A assume eq: "AE x in M. f x = g x"
with borel show "density M f = density M g"
by (auto intro: density_cong)
next
- let ?P = "\<lambda>f A. \<integral>\<^isup>+ x. f x * indicator A x \<partial>M"
+ let ?P = "\<lambda>f A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M"
assume "density M f = density M g"
with borel have eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
by (simp add: emeasure_density[symmetric])
from this[THEN bspec, OF sets.top] fin
- have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
+ have g_fin: "integral\<^sup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
{ fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
- and g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+ and g_fin: "integral\<^sup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
let ?N = "{x\<in>space M. g x < f x}"
have N: "?N \<in> sets M" using borel by simp
- have "?P g ?N \<le> integral\<^isup>P M g" using pos
+ have "?P g ?N \<le> integral\<^sup>P M g" using pos
by (intro positive_integral_mono_AE) (auto split: split_indicator)
then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by auto
- have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
+ have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
by (auto intro!: positive_integral_cong simp: indicator_def)
also have "\<dots> = ?P f ?N - ?P g ?N"
proof (rule positive_integral_diff)
@@ -817,7 +817,7 @@
lemma (in finite_measure) density_unique_finite_measure:
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"
- assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
+ assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. f' x * indicator A x \<partial>M)"
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
shows "AE x in M. f x = f' x"
proof -
@@ -847,13 +847,13 @@
moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"
proof (rule AE_I')
{ fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
- and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+ and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
let ?A = "\<lambda>i. Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
have "(\<Union>i. ?A i) \<in> null_sets M"
proof (rule null_sets_UN)
fix i ::nat have "?A i \<in> sets M"
using borel Q0(1) by auto
- have "?N (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
+ have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
unfolding eq[OF `?A i \<in> sets M`]
by (auto intro!: positive_integral_mono simp: indicator_def)
also have "\<dots> = i * emeasure M (?A i)"
@@ -887,7 +887,7 @@
shows "AE x in M. f x = f' x"
proof -
obtain h where h_borel: "h \<in> borel_measurable M"
- and fin: "integral\<^isup>P M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
+ and fin: "integral\<^sup>P M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
using Ex_finite_integrable_function by auto
then have h_nn: "AE x in M. 0 \<le> h x" by auto
let ?H = "density M h"
@@ -899,21 +899,21 @@
{ fix A assume "A \<in> sets M"
then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
using pos(1) sets.sets_into_space by (force simp: indicator_def)
- then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
+ then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto }
note h_null_sets = this
{ fix A assume "A \<in> sets M"
- have "(\<integral>\<^isup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
+ have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)"
using `A \<in> sets M` h_borel h_nn f f'
by (intro positive_integral_density[symmetric]) auto
- also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)"
by (simp_all add: density_eq)
- also have "\<dots> = (\<integral>\<^isup>+x. f' x * (h x * indicator A x) \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h x * indicator A x) \<partial>M)"
using `A \<in> sets M` h_borel h_nn f f'
by (intro positive_integral_density) auto
- finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)"
+ finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)"
by (simp add: ac_simps)
- then have "(\<integral>\<^isup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^isup>+x. (f' x * indicator A x) \<partial>?H)"
+ then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)"
using `A \<in> sets M` h_borel h_nn f f'
by (subst (asm) (1 2) positive_integral_density[symmetric]) auto }
then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
@@ -950,7 +950,7 @@
fix i
have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))"
unfolding eq ..
- moreover have "(\<integral>\<^isup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"
+ moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"
using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq)
ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x"
using borel pos cover(1) pos
@@ -974,14 +974,14 @@
assume "sigma_finite_measure ?N"
then interpret N: sigma_finite_measure ?N .
from N.Ex_finite_integrable_function obtain h where
- h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<infinity>" and
+ h: "h \<in> borel_measurable M" "integral\<^sup>P ?N h \<noteq> \<infinity>" and
h_nn: "\<And>x. 0 \<le> h x" and
fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
have "AE x in M. f x * h x \<noteq> \<infinity>"
proof (rule AE_I')
- have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn
+ have "integral\<^sup>P ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)" using f h h_nn
by (auto intro!: positive_integral_density)
- then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
+ then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
using h(2) by simp
then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage)
@@ -1030,7 +1030,7 @@
next
fix n obtain i j where
[simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
- have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
+ have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
proof (cases i)
case 0
have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"
@@ -1038,8 +1038,8 @@
from positive_integral_cong_AE[OF this] show ?thesis by simp
next
case (Suc n)
- then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
- (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
+ then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
+ (\<integral>\<^sup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat)
also have "\<dots> = Suc n * emeasure M (Q j)"
using Q by (auto intro!: positive_integral_cmult_indicator)
@@ -1093,11 +1093,11 @@
lemma (in sigma_finite_measure) RN_deriv_positive_integral:
assumes N: "absolutely_continuous M N" "sets N = sets M"
and f: "f \<in> borel_measurable M"
- shows "integral\<^isup>P N f = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
+ shows "integral\<^sup>P N f = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
proof -
- have "integral\<^isup>P N f = integral\<^isup>P (density M (RN_deriv M N)) f"
+ have "integral\<^sup>P N f = integral\<^sup>P (density M (RN_deriv M N)) f"
using N by (simp add: density_RN_deriv)
- also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density)
finally show ?thesis by simp
qed
@@ -1193,7 +1193,7 @@
and f: "f \<in> borel_measurable M"
shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
integrable M (\<lambda>x. real (RN_deriv M N x) * f x)" (is ?integrable)
- and RN_deriv_integral: "integral\<^isup>L N f =
+ and RN_deriv_integral: "integral\<^sup>L N f =
(\<integral>x. real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
proof -
note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
@@ -1207,8 +1207,8 @@
by (simp add: mult_le_0_iff)
then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"
using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) }
- then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M N x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (f x) \<partial>M)"
- "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M N x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (- f x) \<partial>M)"
+ then have "(\<integral>\<^sup>+x. ereal (real (RN_deriv M N x) * f x) \<partial>M) = (\<integral>\<^sup>+x. RN_deriv M N x * ereal (f x) \<partial>M)"
+ "(\<integral>\<^sup>+x. ereal (- (real (RN_deriv M N x) * f x)) \<partial>M) = (\<integral>\<^sup>+x. RN_deriv M N x * ereal (- f x) \<partial>M)"
using RN_deriv_finite[OF N ac] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
by (auto intro!: positive_integral_cong_AE) }
note * = this
@@ -1235,9 +1235,9 @@
show "(\<lambda>x. real (RN_deriv M N x)) \<in> borel_measurable M"
using RN by auto
- have "N (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
+ have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
- also have "\<dots> = (\<integral>\<^isup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
by (intro positive_integral_cong) (auto simp: indicator_def)
also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
using RN by (intro positive_integral_cmult_indicator) auto
@@ -1261,9 +1261,9 @@
show "\<And>x. 0 \<le> real (RN_deriv M N x)"
using RN by (auto intro: real_of_ereal_pos)
- have "N (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
+ have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
- also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)"
+ also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)"
by (intro positive_integral_cong) (auto simp: indicator_def)
finally have "AE x in N. RN_deriv M N x \<noteq> 0"
using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
@@ -1278,7 +1278,7 @@
proof -
note deriv = RN_deriv[OF ac]
from deriv(1,3) `{x} \<in> sets M`
- have "density M (RN_deriv M N) {x} = (\<integral>\<^isup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
+ have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong)
with x deriv show ?thesis
by (auto simp: positive_integral_cmult_indicator)
--- a/src/HOL/Quickcheck_Examples/Hotel_Example.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Quickcheck_Examples/Hotel_Example.thy Tue Aug 13 16:25:47 2013 +0200
@@ -76,9 +76,9 @@
definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
where
- "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
- s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
- no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
+ "feels_safe s r = (\<exists>s\<^sub>1 s\<^sub>2 s\<^sub>3 g c c'.
+ s = s\<^sub>3 @ [Enter g r c] @ s\<^sub>2 @ [Check_in g r c'] @ s\<^sub>1 \<and>
+ no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})"
section {* Some setup *}
--- a/src/HOL/Quickcheck_Exhaustive.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Quickcheck_Exhaustive.thy Tue Aug 13 16:25:47 2013 +0200
@@ -447,11 +447,11 @@
begin
definition
- "check_all f = f (Code_Evaluation.valtermify Enum.finite_1.a\<^isub>1)"
+ "check_all f = f (Code_Evaluation.valtermify Enum.finite_1.a\<^sub>1)"
definition enum_term_of_finite_1 :: "Enum.finite_1 itself => unit => term list"
where
- "enum_term_of_finite_1 = (%_ _. [Code_Evaluation.term_of Enum.finite_1.a\<^isub>1])"
+ "enum_term_of_finite_1 = (%_ _. [Code_Evaluation.term_of Enum.finite_1.a\<^sub>1])"
instance ..
@@ -461,8 +461,8 @@
begin
definition
- "check_all f = (f (Code_Evaluation.valtermify Enum.finite_2.a\<^isub>1)
- orelse f (Code_Evaluation.valtermify Enum.finite_2.a\<^isub>2))"
+ "check_all f = (f (Code_Evaluation.valtermify Enum.finite_2.a\<^sub>1)
+ orelse f (Code_Evaluation.valtermify Enum.finite_2.a\<^sub>2))"
definition enum_term_of_finite_2 :: "Enum.finite_2 itself => unit => term list"
where
@@ -476,9 +476,9 @@
begin
definition
- "check_all f = (f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>1)
- orelse f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>2)
- orelse f (Code_Evaluation.valtermify Enum.finite_3.a\<^isub>3))"
+ "check_all f = (f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>1)
+ orelse f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>2)
+ orelse f (Code_Evaluation.valtermify Enum.finite_3.a\<^sub>3))"
definition enum_term_of_finite_3 :: "Enum.finite_3 itself => unit => term list"
where
@@ -492,10 +492,10 @@
begin
definition
- "check_all f = (f (Code_Evaluation.valtermify Enum.finite_4.a\<^isub>1)
- orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^isub>2)
- orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^isub>3)
- orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^isub>4))"
+ "check_all f = (f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>1)
+ orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>2)
+ orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>3)
+ orelse f (Code_Evaluation.valtermify Enum.finite_4.a\<^sub>4))"
definition enum_term_of_finite_4 :: "Enum.finite_4 itself => unit => term list"
where
--- a/src/HOL/Rat.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Rat.thy Tue Aug 13 16:25:47 2013 +0200
@@ -409,9 +409,9 @@
assume "b \<noteq> 0" and "d \<noteq> 0" and "a * d = c * b"
hence "a * d * b * d = c * b * b * d"
by simp
- hence "a * b * d\<twosuperior> = c * d * b\<twosuperior>"
+ hence "a * b * d\<^sup>2 = c * d * b\<^sup>2"
unfolding power2_eq_square by (simp add: mult_ac)
- hence "0 < a * b * d\<twosuperior> \<longleftrightarrow> 0 < c * d * b\<twosuperior>"
+ hence "0 < a * b * d\<^sup>2 \<longleftrightarrow> 0 < c * d * b\<^sup>2"
by simp
thus "0 < a * b \<longleftrightarrow> 0 < c * d"
using `b \<noteq> 0` and `d \<noteq> 0`
--- a/src/HOL/Statespace/DistinctTreeProver.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Statespace/DistinctTreeProver.thy Tue Aug 13 16:25:47 2013 +0200
@@ -303,29 +303,29 @@
| None \<Rightarrow> None)"
lemma subtract_Some_set_of_res:
- "subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^isub>2"
-proof (induct t\<^isub>1 arbitrary: t\<^isub>2 t)
+ "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^sub>2"
+proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
case Tip thus ?case by simp
next
case (Node l x b r)
- have sub: "subtract (Node l x b r) t\<^isub>2 = Some t" by fact
+ have sub: "subtract (Node l x b r) t\<^sub>2 = Some t" by fact
show ?case
- proof (cases "delete x t\<^isub>2")
- case (Some t\<^isub>2')
+ proof (cases "delete x t\<^sub>2")
+ case (Some t\<^sub>2')
note del_x_Some = this
from delete_Some_set_of [OF Some]
- have t2'_t2: "set_of t\<^isub>2' \<subseteq> set_of t\<^isub>2" .
+ have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
show ?thesis
- proof (cases "subtract l t\<^isub>2'")
- case (Some t\<^isub>2'')
+ proof (cases "subtract l t\<^sub>2'")
+ case (Some t\<^sub>2'')
note sub_l_Some = this
from Node.hyps (1) [OF Some]
- have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
+ have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
show ?thesis
- proof (cases "subtract r t\<^isub>2''")
- case (Some t\<^isub>2''')
+ proof (cases "subtract r t\<^sub>2''")
+ case (Some t\<^sub>2''')
from Node.hyps (2) [OF Some ]
- have "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''" .
+ have "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''" .
with Some sub_l_Some del_x_Some sub t2''_t2' t2'_t2
show ?thesis
by simp
@@ -348,35 +348,35 @@
qed
lemma subtract_Some_set_of:
- "subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> set_of t\<^isub>1 \<subseteq> set_of t\<^isub>2"
-proof (induct t\<^isub>1 arbitrary: t\<^isub>2 t)
+ "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<subseteq> set_of t\<^sub>2"
+proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
case Tip thus ?case by simp
next
case (Node l x d r)
- have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
+ have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
show ?case
- proof (cases "delete x t\<^isub>2")
- case (Some t\<^isub>2')
+ proof (cases "delete x t\<^sub>2")
+ case (Some t\<^sub>2')
note del_x_Some = this
from delete_Some_set_of [OF Some]
- have t2'_t2: "set_of t\<^isub>2' \<subseteq> set_of t\<^isub>2" .
- from delete_None_set_of_conv [of x t\<^isub>2] Some
- have x_t2: "x \<in> set_of t\<^isub>2"
+ have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
+ from delete_None_set_of_conv [of x t\<^sub>2] Some
+ have x_t2: "x \<in> set_of t\<^sub>2"
by simp
show ?thesis
- proof (cases "subtract l t\<^isub>2'")
- case (Some t\<^isub>2'')
+ proof (cases "subtract l t\<^sub>2'")
+ case (Some t\<^sub>2'')
note sub_l_Some = this
from Node.hyps (1) [OF Some]
- have l_t2': "set_of l \<subseteq> set_of t\<^isub>2'" .
+ have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" .
from subtract_Some_set_of_res [OF Some]
- have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
+ have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
show ?thesis
- proof (cases "subtract r t\<^isub>2''")
- case (Some t\<^isub>2''')
+ proof (cases "subtract r t\<^sub>2''")
+ case (Some t\<^sub>2''')
from Node.hyps (2) [OF Some ]
- have r_t\<^isub>2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
- from Some sub_l_Some del_x_Some sub r_t\<^isub>2'' l_t2' t2'_t2 t2''_t2' x_t2
+ have r_t\<^sub>2'': "set_of r \<subseteq> set_of t\<^sub>2''" .
+ from Some sub_l_Some del_x_Some sub r_t\<^sub>2'' l_t2' t2'_t2 t2''_t2' x_t2
show ?thesis
by auto
next
@@ -398,30 +398,30 @@
qed
lemma subtract_Some_all_distinct_res:
- "subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> all_distinct t\<^isub>2 \<Longrightarrow> all_distinct t"
-proof (induct t\<^isub>1 arbitrary: t\<^isub>2 t)
+ "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t"
+proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
case Tip thus ?case by simp
next
case (Node l x d r)
- have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
- have dist_t2: "all_distinct t\<^isub>2" by fact
+ have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
+ have dist_t2: "all_distinct t\<^sub>2" by fact
show ?case
- proof (cases "delete x t\<^isub>2")
- case (Some t\<^isub>2')
+ proof (cases "delete x t\<^sub>2")
+ case (Some t\<^sub>2')
note del_x_Some = this
from delete_Some_all_distinct [OF Some dist_t2]
- have dist_t2': "all_distinct t\<^isub>2'" .
+ have dist_t2': "all_distinct t\<^sub>2'" .
show ?thesis
- proof (cases "subtract l t\<^isub>2'")
- case (Some t\<^isub>2'')
+ proof (cases "subtract l t\<^sub>2'")
+ case (Some t\<^sub>2'')
note sub_l_Some = this
from Node.hyps (1) [OF Some dist_t2']
- have dist_t2'': "all_distinct t\<^isub>2''" .
+ have dist_t2'': "all_distinct t\<^sub>2''" .
show ?thesis
- proof (cases "subtract r t\<^isub>2''")
- case (Some t\<^isub>2''')
+ proof (cases "subtract r t\<^sub>2''")
+ case (Some t\<^sub>2''')
from Node.hyps (2) [OF Some dist_t2'']
- have dist_t2''': "all_distinct t\<^isub>2'''" .
+ have dist_t2''': "all_distinct t\<^sub>2'''" .
from Some sub_l_Some del_x_Some sub
dist_t2'''
show ?thesis
@@ -446,36 +446,36 @@
lemma subtract_Some_dist_res:
- "subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> set_of t\<^isub>1 \<inter> set_of t = {}"
-proof (induct t\<^isub>1 arbitrary: t\<^isub>2 t)
+ "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<inter> set_of t = {}"
+proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
case Tip thus ?case by simp
next
case (Node l x d r)
- have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
+ have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
show ?case
- proof (cases "delete x t\<^isub>2")
- case (Some t\<^isub>2')
+ proof (cases "delete x t\<^sub>2")
+ case (Some t\<^sub>2')
note del_x_Some = this
from delete_Some_x_set_of [OF Some]
- obtain x_t2: "x \<in> set_of t\<^isub>2" and x_not_t2': "x \<notin> set_of t\<^isub>2'"
+ obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'"
by simp
from delete_Some_set_of [OF Some]
- have t2'_t2: "set_of t\<^isub>2' \<subseteq> set_of t\<^isub>2" .
+ have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
show ?thesis
- proof (cases "subtract l t\<^isub>2'")
- case (Some t\<^isub>2'')
+ proof (cases "subtract l t\<^sub>2'")
+ case (Some t\<^sub>2'')
note sub_l_Some = this
from Node.hyps (1) [OF Some ]
- have dist_l_t2'': "set_of l \<inter> set_of t\<^isub>2'' = {}".
+ have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
from subtract_Some_set_of_res [OF Some]
- have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
+ have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
show ?thesis
- proof (cases "subtract r t\<^isub>2''")
- case (Some t\<^isub>2''')
+ proof (cases "subtract r t\<^sub>2''")
+ case (Some t\<^sub>2''')
from Node.hyps (2) [OF Some]
- have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}" .
+ have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}" .
from subtract_Some_set_of_res [OF Some]
- have t2'''_t2'': "set_of t\<^isub>2''' \<subseteq> set_of t\<^isub>2''".
+ have t2'''_t2'': "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''".
from Some sub_l_Some del_x_Some sub t2'''_t2'' dist_l_t2'' dist_r_t2'''
t2''_t2' t2'_t2 x_not_t2'
@@ -500,48 +500,48 @@
qed
lemma subtract_Some_all_distinct:
- "subtract t\<^isub>1 t\<^isub>2 = Some t \<Longrightarrow> all_distinct t\<^isub>2 \<Longrightarrow> all_distinct t\<^isub>1"
-proof (induct t\<^isub>1 arbitrary: t\<^isub>2 t)
+ "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t\<^sub>1"
+proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t)
case Tip thus ?case by simp
next
case (Node l x d r)
- have sub: "subtract (Node l x d r) t\<^isub>2 = Some t" by fact
- have dist_t2: "all_distinct t\<^isub>2" by fact
+ have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact
+ have dist_t2: "all_distinct t\<^sub>2" by fact
show ?case
- proof (cases "delete x t\<^isub>2")
- case (Some t\<^isub>2')
+ proof (cases "delete x t\<^sub>2")
+ case (Some t\<^sub>2')
note del_x_Some = this
from delete_Some_all_distinct [OF Some dist_t2 ]
- have dist_t2': "all_distinct t\<^isub>2'" .
+ have dist_t2': "all_distinct t\<^sub>2'" .
from delete_Some_set_of [OF Some]
- have t2'_t2: "set_of t\<^isub>2' \<subseteq> set_of t\<^isub>2" .
+ have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" .
from delete_Some_x_set_of [OF Some]
- obtain x_t2: "x \<in> set_of t\<^isub>2" and x_not_t2': "x \<notin> set_of t\<^isub>2'"
+ obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'"
by simp
show ?thesis
- proof (cases "subtract l t\<^isub>2'")
- case (Some t\<^isub>2'')
+ proof (cases "subtract l t\<^sub>2'")
+ case (Some t\<^sub>2'')
note sub_l_Some = this
from Node.hyps (1) [OF Some dist_t2' ]
have dist_l: "all_distinct l" .
from subtract_Some_all_distinct_res [OF Some dist_t2']
- have dist_t2'': "all_distinct t\<^isub>2''" .
+ have dist_t2'': "all_distinct t\<^sub>2''" .
from subtract_Some_set_of [OF Some]
- have l_t2': "set_of l \<subseteq> set_of t\<^isub>2'" .
+ have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" .
from subtract_Some_set_of_res [OF Some]
- have t2''_t2': "set_of t\<^isub>2'' \<subseteq> set_of t\<^isub>2'" .
+ have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" .
from subtract_Some_dist_res [OF Some]
- have dist_l_t2'': "set_of l \<inter> set_of t\<^isub>2'' = {}".
+ have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
show ?thesis
- proof (cases "subtract r t\<^isub>2''")
- case (Some t\<^isub>2''')
+ proof (cases "subtract r t\<^sub>2''")
+ case (Some t\<^sub>2''')
from Node.hyps (2) [OF Some dist_t2'']
have dist_r: "all_distinct r" .
from subtract_Some_set_of [OF Some]
- have r_t2'': "set_of r \<subseteq> set_of t\<^isub>2''" .
+ have r_t2'': "set_of r \<subseteq> set_of t\<^sub>2''" .
from subtract_Some_dist_res [OF Some]
- have dist_r_t2''': "set_of r \<inter> set_of t\<^isub>2''' = {}".
+ have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}".
from dist_l dist_r Some sub_l_Some del_x_Some r_t2'' l_t2' x_t2 x_not_t2'
t2''_t2' dist_l_t2'' dist_r_t2'''
--- a/src/HOL/Tools/ATP/atp_util.ML Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Tools/ATP/atp_util.ML Tue Aug 13 16:25:47 2013 +0200
@@ -130,7 +130,7 @@
| NONE => [])
| NONE => []
-val subscript = implode o map (prefix "\<^isub>") o raw_explode (* FIXME Symbol.explode (?) *)
+val subscript = implode o map (prefix "\<^sub>") o raw_explode (* FIXME Symbol.explode (?) *)
fun nat_subscript n =
n |> string_of_int |> print_mode_active Symbol.xsymbolsN ? subscript
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Tue Aug 13 16:25:47 2013 +0200
@@ -125,7 +125,7 @@
fun atom_suffix s =
nat_subscript
#> (s <> "" andalso Symbol.is_ascii_digit (List.last (raw_explode s))) (* FIXME Symbol.explode (?) *)
- ? prefix "\<^isub>,"
+ ? prefix "\<^sub>,"
fun nth_atom_name thy atomss pool prefix T j =
let
val ss = these (triple_lookup (type_match thy) atomss T)
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Tue Aug 13 16:25:47 2013 +0200
@@ -238,7 +238,7 @@
val parse_time_option = Sledgehammer_Util.parse_time_option
val string_of_time = ATP_Util.string_of_time
-val i_subscript = implode o map (prefix "\<^isub>") o Symbol.explode
+val i_subscript = implode o map (prefix "\<^sub>") o Symbol.explode
fun nat_subscript n =
n |> signed_string_of_int |> print_mode_active Symbol.xsymbolsN ? i_subscript
--- a/src/HOL/Transcendental.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Transcendental.thy Tue Aug 13 16:25:47 2013 +0200
@@ -1266,13 +1266,13 @@
by (rule le_imp_inverse_le) simp
hence "inverse (fact (n + 2)) \<le> 1/2 * (1/2)^n"
by (simp add: inverse_mult_distrib power_inverse)
- hence "inverse (fact (n + 2)) * (x^n * x\<twosuperior>) \<le> 1/2 * (1/2)^n * (1 * x\<twosuperior>)"
+ hence "inverse (fact (n + 2)) * (x^n * x\<^sup>2) \<le> 1/2 * (1/2)^n * (1 * x\<^sup>2)"
by (rule mult_mono)
(rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg)
- hence "inverse (fact (n + 2)) * x ^ (n + 2) \<le> (x\<twosuperior>/2) * ((1/2)^n)"
+ hence "inverse (fact (n + 2)) * x ^ (n + 2) \<le> (x\<^sup>2/2) * ((1/2)^n)"
unfolding power_add by (simp add: mult_ac del: fact_Suc) }
note aux1 = this
- have "(\<lambda>n. x\<twosuperior> / 2 * (1 / 2) ^ n) sums (x\<twosuperior> / 2 * (1 / (1 - 1 / 2)))"
+ have "(\<lambda>n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1 / 2)))"
by (intro sums_mult geometric_sums, simp)
hence aux2: "(\<lambda>n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2"
by simp
@@ -1299,7 +1299,7 @@
also have "... <= 1"
by (auto simp add: a)
finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
- moreover have c: "0 < 1 + x + x\<twosuperior>"
+ moreover have c: "0 < 1 + x + x\<^sup>2"
by (simp add: add_pos_nonneg a)
ultimately have "1 - x <= 1 / (1 + x + x^2)"
by (elim mult_imp_le_div_pos)
@@ -2134,25 +2134,25 @@
lemma cos_zero [simp]: "cos 0 = 1"
unfolding cos_def cos_coeff_def by (simp add: powser_zero)
-lemma sin_cos_squared_add [simp]: "(sin x)\<twosuperior> + (cos x)\<twosuperior> = 1"
+lemma sin_cos_squared_add [simp]: "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1"
proof -
- have "\<forall>x. DERIV (\<lambda>x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0"
+ have "\<forall>x. DERIV (\<lambda>x. (sin x)\<^sup>2 + (cos x)\<^sup>2) x :> 0"
by (auto intro!: DERIV_intros)
- hence "(sin x)\<twosuperior> + (cos x)\<twosuperior> = (sin 0)\<twosuperior> + (cos 0)\<twosuperior>"
+ hence "(sin x)\<^sup>2 + (cos x)\<^sup>2 = (sin 0)\<^sup>2 + (cos 0)\<^sup>2"
by (rule DERIV_isconst_all)
- thus "(sin x)\<twosuperior> + (cos x)\<twosuperior> = 1" by simp
+ thus "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1" by simp
qed
-lemma sin_cos_squared_add2 [simp]: "(cos x)\<twosuperior> + (sin x)\<twosuperior> = 1"
+lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1"
by (subst add_commute, rule sin_cos_squared_add)
lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1"
using sin_cos_squared_add2 [unfolded power2_eq_square] .
-lemma sin_squared_eq: "(sin x)\<twosuperior> = 1 - (cos x)\<twosuperior>"
+lemma sin_squared_eq: "(sin x)\<^sup>2 = 1 - (cos x)\<^sup>2"
unfolding eq_diff_eq by (rule sin_cos_squared_add)
-lemma cos_squared_eq: "(cos x)\<twosuperior> = 1 - (sin x)\<twosuperior>"
+lemma cos_squared_eq: "(cos x)\<^sup>2 = 1 - (sin x)\<^sup>2"
unfolding eq_diff_eq by (rule sin_cos_squared_add2)
lemma abs_sin_le_one [simp]: "\<bar>sin x\<bar> \<le> 1"
@@ -2208,7 +2208,7 @@
using sin_cos_add_lemma unfolding realpow_two_sum_zero_iff by simp
lemma sin_cos_minus_lemma:
- "(sin(-x) + sin(x))\<twosuperior> + (cos(-x) - cos(x))\<twosuperior> = 0" (is "?f x = 0")
+ "(sin(-x) + sin(x))\<^sup>2 + (cos(-x) - cos(x))\<^sup>2 = 0" (is "?f x = 0")
proof -
have "\<forall>x. DERIV (\<lambda>x. ?f x) x :> 0"
by (auto intro!: DERIV_intros simp add: algebra_simps)
@@ -2238,7 +2238,7 @@
lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x"
using sin_add [where x=x and y=x] by simp
-lemma cos_double: "cos(2* x) = ((cos x)\<twosuperior>) - ((sin x)\<twosuperior>)"
+lemma cos_double: "cos(2* x) = ((cos x)\<^sup>2) - ((sin x)\<^sup>2)"
using cos_add [where x=x and y=x]
by (simp add: power2_eq_square)
@@ -2732,7 +2732,7 @@
unfolding tan_def sin_double cos_double sin_squared_eq
by (simp add: power2_eq_square)
-lemma DERIV_tan [simp]: "cos x \<noteq> 0 \<Longrightarrow> DERIV tan x :> inverse ((cos x)\<twosuperior>)"
+lemma DERIV_tan [simp]: "cos x \<noteq> 0 \<Longrightarrow> DERIV tan x :> inverse ((cos x)\<^sup>2)"
unfolding tan_def
by (auto intro!: DERIV_intros, simp add: divide_inverse power2_eq_square)
@@ -2827,10 +2827,10 @@
thus "DERIV tan x' :> inverse (cos x'^2)" by (rule DERIV_tan)
qed
from MVT2[OF `y < x` this]
- obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<twosuperior>)" by auto
+ obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto
hence "- (pi / 2) < z" and "z < pi / 2" using assms by auto
hence "0 < cos z" using cos_gt_zero_pi by auto
- hence inv_pos: "0 < inverse ((cos z)\<twosuperior>)" by auto
+ hence inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto
have "0 < x - y" using `y < x` by auto
from mult_pos_pos [OF this inv_pos]
have "0 < tan x - tan y" unfolding tan_diff by auto
@@ -2978,27 +2978,27 @@
apply (auto intro!: the1_equality cos_total)
done
-lemma cos_arcsin: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cos (arcsin x) = sqrt (1 - x\<twosuperior>)"
-apply (subgoal_tac "x\<twosuperior> \<le> 1")
+lemma cos_arcsin: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cos (arcsin x) = sqrt (1 - x\<^sup>2)"
+apply (subgoal_tac "x\<^sup>2 \<le> 1")
apply (rule power2_eq_imp_eq)
apply (simp add: cos_squared_eq)
apply (rule cos_ge_zero)
apply (erule (1) arcsin_lbound)
apply (erule (1) arcsin_ubound)
apply simp
-apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> \<le> 1\<twosuperior>", simp)
+apply (subgoal_tac "\<bar>x\<bar>\<^sup>2 \<le> 1\<^sup>2", simp)
apply (rule power_mono, simp, simp)
done
-lemma sin_arccos: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> sin (arccos x) = sqrt (1 - x\<twosuperior>)"
-apply (subgoal_tac "x\<twosuperior> \<le> 1")
+lemma sin_arccos: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> sin (arccos x) = sqrt (1 - x\<^sup>2)"
+apply (subgoal_tac "x\<^sup>2 \<le> 1")
apply (rule power2_eq_imp_eq)
apply (simp add: sin_squared_eq)
apply (rule sin_ge_zero)
apply (erule (1) arccos_lbound)
apply (erule (1) arccos_ubound)
apply simp
-apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> \<le> 1\<twosuperior>", simp)
+apply (subgoal_tac "\<bar>x\<bar>\<^sup>2 \<le> 1\<^sup>2", simp)
apply (rule power_mono, simp, simp)
done
@@ -3041,26 +3041,26 @@
by (intro less_imp_neq [symmetric] cos_gt_zero_pi
arctan_lbound arctan_ubound)
-lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<twosuperior>)"
+lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<^sup>2)"
proof (rule power2_eq_imp_eq)
- have "0 < 1 + x\<twosuperior>" by (simp add: add_pos_nonneg)
- show "0 \<le> 1 / sqrt (1 + x\<twosuperior>)" by simp
+ have "0 < 1 + x\<^sup>2" by (simp add: add_pos_nonneg)
+ show "0 \<le> 1 / sqrt (1 + x\<^sup>2)" by simp
show "0 \<le> cos (arctan x)"
by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound)
- have "(cos (arctan x))\<twosuperior> * (1 + (tan (arctan x))\<twosuperior>) = 1"
+ have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1"
unfolding tan_def by (simp add: distrib_left power_divide)
- thus "(cos (arctan x))\<twosuperior> = (1 / sqrt (1 + x\<twosuperior>))\<twosuperior>"
- using `0 < 1 + x\<twosuperior>` by (simp add: power_divide eq_divide_eq)
+ thus "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2"
+ using `0 < 1 + x\<^sup>2` by (simp add: power_divide eq_divide_eq)
qed
-lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<twosuperior>)"
+lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)"
using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]]
using tan_arctan [of x] unfolding tan_def cos_arctan
by (simp add: eq_divide_eq)
lemma tan_sec: "cos x \<noteq> 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2"
apply (rule power_inverse [THEN subst])
-apply (rule_tac c1 = "(cos x)\<twosuperior>" in real_mult_right_cancel [THEN iffD1])
+apply (rule_tac c1 = "(cos x)\<^sup>2" in real_mult_right_cancel [THEN iffD1])
apply (auto dest: field_power_not_zero
simp add: power_mult_distrib distrib_right power_divide tan_def
mult_assoc power_inverse [symmetric])
@@ -3151,11 +3151,11 @@
unfolding continuous_on_def by (auto intro: tendsto_arctan)
lemma DERIV_arcsin:
- "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arcsin x :> inverse (sqrt (1 - x\<twosuperior>))"
+ "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arcsin x :> inverse (sqrt (1 - x\<^sup>2))"
apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"])
apply (rule DERIV_cong [OF DERIV_sin])
apply (simp add: cos_arcsin)
-apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> < 1\<twosuperior>", simp)
+apply (subgoal_tac "\<bar>x\<bar>\<^sup>2 < 1\<^sup>2", simp)
apply (rule power_strict_mono, simp, simp, simp)
apply assumption
apply assumption
@@ -3164,11 +3164,11 @@
done
lemma DERIV_arccos:
- "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arccos x :> inverse (- sqrt (1 - x\<twosuperior>))"
+ "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arccos x :> inverse (- sqrt (1 - x\<^sup>2))"
apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"])
apply (rule DERIV_cong [OF DERIV_cos])
apply (simp add: sin_arccos)
-apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> < 1\<twosuperior>", simp)
+apply (subgoal_tac "\<bar>x\<bar>\<^sup>2 < 1\<^sup>2", simp)
apply (rule power_strict_mono, simp, simp, simp)
apply assumption
apply assumption
@@ -3176,12 +3176,12 @@
apply (erule (1) isCont_arccos)
done
-lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<twosuperior>)"
+lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)"
apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"])
apply (rule DERIV_cong [OF DERIV_tan])
apply (rule cos_arctan_not_zero)
apply (simp add: power_inverse tan_sec [symmetric])
-apply (subgoal_tac "0 < 1 + x\<twosuperior>", simp)
+apply (subgoal_tac "0 < 1 + x\<^sup>2", simp)
apply (simp add: add_pos_nonneg)
apply (simp, simp, simp, rule isCont_arctan)
done
@@ -3233,11 +3233,11 @@
by (simp add: cos_ge_zero)
have "0 = cos (pi / 4 + pi / 4)"
by simp
- also have "cos (pi / 4 + pi / 4) = ?c\<twosuperior> - ?s\<twosuperior>"
+ also have "cos (pi / 4 + pi / 4) = ?c\<^sup>2 - ?s\<^sup>2"
by (simp only: cos_add power2_eq_square)
- also have "\<dots> = 2 * ?c\<twosuperior> - 1"
+ also have "\<dots> = 2 * ?c\<^sup>2 - 1"
by (simp add: sin_squared_eq)
- finally have "?c\<twosuperior> = (sqrt 2 / 2)\<twosuperior>"
+ finally have "?c\<^sup>2 = (sqrt 2 / 2)\<^sup>2"
by (simp add: power_divide)
thus ?thesis
using nonneg by (rule power2_eq_imp_eq) simp
@@ -3252,9 +3252,9 @@
by simp
also have "\<dots> = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s"
by (simp only: cos_add sin_add)
- also have "\<dots> = ?c * (?c\<twosuperior> - 3 * ?s\<twosuperior>)"
+ also have "\<dots> = ?c * (?c\<^sup>2 - 3 * ?s\<^sup>2)"
by (simp add: algebra_simps power2_eq_square)
- finally have "?c\<twosuperior> = (sqrt 3 / 2)\<twosuperior>"
+ finally have "?c\<^sup>2 = (sqrt 3 / 2)\<^sup>2"
using pos_c by (simp add: sin_squared_eq power_divide)
thus ?thesis
using pos_c [THEN order_less_imp_le]
@@ -3695,7 +3695,7 @@
subsection {* Existence of Polar Coordinates *}
-lemma cos_x_y_le_one: "\<bar>x / sqrt (x\<twosuperior> + y\<twosuperior>)\<bar> \<le> 1"
+lemma cos_x_y_le_one: "\<bar>x / sqrt (x\<^sup>2 + y\<^sup>2)\<bar> \<le> 1"
apply (rule power2_le_imp_le [OF _ zero_le_one])
apply (simp add: power_divide divide_le_eq not_sum_power2_lt_zero)
done
@@ -3703,7 +3703,7 @@
lemma cos_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> cos (arccos y) = y"
by (simp add: abs_le_iff)
-lemma sin_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> sin (arccos y) = sqrt (1 - y\<twosuperior>)"
+lemma sin_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> sin (arccos y) = sqrt (1 - y\<^sup>2)"
by (simp add: sin_arccos abs_le_iff)
lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one]
@@ -3712,8 +3712,8 @@
lemma polar_ex1:
"0 < y ==> \<exists>r a. x = r * cos a & y = r * sin a"
-apply (rule_tac x = "sqrt (x\<twosuperior> + y\<twosuperior>)" in exI)
-apply (rule_tac x = "arccos (x / sqrt (x\<twosuperior> + y\<twosuperior>))" in exI)
+apply (rule_tac x = "sqrt (x\<^sup>2 + y\<^sup>2)" in exI)
+apply (rule_tac x = "arccos (x / sqrt (x\<^sup>2 + y\<^sup>2))" in exI)
apply (simp add: cos_arccos_lemma1)
apply (simp add: sin_arccos_lemma1)
apply (simp add: power_divide)
--- a/src/HOL/Unix/Unix.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/Unix/Unix.thy Tue Aug 13 16:25:47 2013 +0200
@@ -712,27 +712,27 @@
apply (rule can_exec_nil)
done
-theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^isub>1 \<Longrightarrow>
- Readable \<in> perms\<^isub>2 \<Longrightarrow> name\<^isub>3 \<noteq> name\<^isub>4 \<Longrightarrow>
+theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^sub>1 \<Longrightarrow>
+ Readable \<in> perms\<^sub>2 \<Longrightarrow> name\<^sub>3 \<noteq> name\<^sub>4 \<Longrightarrow>
can_exec (init users)
- [Mkdir u perms\<^isub>1 [u, name\<^isub>1],
- Mkdir u' perms\<^isub>2 [u, name\<^isub>1, name\<^isub>2],
- Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
- Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>4],
- Readdir u {name\<^isub>3, name\<^isub>4} [u, name\<^isub>1, name\<^isub>2]]"
+ [Mkdir u perms\<^sub>1 [u, name\<^sub>1],
+ Mkdir u' perms\<^sub>2 [u, name\<^sub>1, name\<^sub>2],
+ Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>3],
+ Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>4],
+ Readdir u {name\<^sub>3, name\<^sub>4} [u, name\<^sub>1, name\<^sub>2]]"
apply (rule can_exec_cons, rule transition.intros,
(force simp add: eval)+, (simp add: eval)?)+
txt {* peek at result: @{subgoals [display]} *}
apply (rule can_exec_nil)
done
-theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^isub>1 \<Longrightarrow> Readable \<in> perms\<^isub>3 \<Longrightarrow>
+theorem "u \<in> users \<Longrightarrow> Writable \<in> perms\<^sub>1 \<Longrightarrow> Readable \<in> perms\<^sub>3 \<Longrightarrow>
can_exec (init users)
- [Mkdir u perms\<^isub>1 [u, name\<^isub>1],
- Mkdir u' perms\<^isub>2 [u, name\<^isub>1, name\<^isub>2],
- Creat u' perms\<^isub>3 [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
- Write u' ''foo'' [u, name\<^isub>1, name\<^isub>2, name\<^isub>3],
- Read u ''foo'' [u, name\<^isub>1, name\<^isub>2, name\<^isub>3]]"
+ [Mkdir u perms\<^sub>1 [u, name\<^sub>1],
+ Mkdir u' perms\<^sub>2 [u, name\<^sub>1, name\<^sub>2],
+ Creat u' perms\<^sub>3 [u, name\<^sub>1, name\<^sub>2, name\<^sub>3],
+ Write u' ''foo'' [u, name\<^sub>1, name\<^sub>2, name\<^sub>3],
+ Read u ''foo'' [u, name\<^sub>1, name\<^sub>2, name\<^sub>3]]"
apply (rule can_exec_cons, rule transition.intros,
(force simp add: eval)+, (simp add: eval)?)+
txt {* peek at result: @{subgoals [display]} *}
@@ -784,7 +784,7 @@
Here @{prop "P x"} refers to the restriction on file-system
operations that are admitted after having reached the critical
configuration; according to the problem specification this will
- become @{prop "uid_of x = user\<^isub>1"} later on. Furthermore,
+ become @{prop "uid_of x = user\<^sub>1"} later on. Furthermore,
@{term y} refers to the operations we claim to be impossible to
perform afterwards, we will take @{term Rmdir} later. Moreover
@{term Q} is a suitable (auxiliary) invariant over the file-system;
@@ -803,29 +803,29 @@
locale situation =
fixes users :: "uid set"
- and user\<^isub>1 :: uid
- and user\<^isub>2 :: uid
- and name\<^isub>1 :: name
- and name\<^isub>2 :: name
- and name\<^isub>3 :: name
- and perms\<^isub>1 :: perms
- and perms\<^isub>2 :: perms
- assumes user\<^isub>1_known: "user\<^isub>1 \<in> users"
- and user\<^isub>1_not_root: "user\<^isub>1 \<noteq> 0"
- and users_neq: "user\<^isub>1 \<noteq> user\<^isub>2"
- and perms\<^isub>1_writable: "Writable \<in> perms\<^isub>1"
- and perms\<^isub>2_not_writable: "Writable \<notin> perms\<^isub>2"
- notes facts = user\<^isub>1_known user\<^isub>1_not_root users_neq
- perms\<^isub>1_writable perms\<^isub>2_not_writable
+ and user\<^sub>1 :: uid
+ and user\<^sub>2 :: uid
+ and name\<^sub>1 :: name
+ and name\<^sub>2 :: name
+ and name\<^sub>3 :: name
+ and perms\<^sub>1 :: perms
+ and perms\<^sub>2 :: perms
+ assumes user\<^sub>1_known: "user\<^sub>1 \<in> users"
+ and user\<^sub>1_not_root: "user\<^sub>1 \<noteq> 0"
+ and users_neq: "user\<^sub>1 \<noteq> user\<^sub>2"
+ and perms\<^sub>1_writable: "Writable \<in> perms\<^sub>1"
+ and perms\<^sub>2_not_writable: "Writable \<notin> perms\<^sub>2"
+ notes facts = user\<^sub>1_known user\<^sub>1_not_root users_neq
+ perms\<^sub>1_writable perms\<^sub>2_not_writable
begin
definition
"bogus =
- [Mkdir user\<^isub>1 perms\<^isub>1 [user\<^isub>1, name\<^isub>1],
- Mkdir user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2],
- Creat user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2, name\<^isub>3]]"
+ [Mkdir user\<^sub>1 perms\<^sub>1 [user\<^sub>1, name\<^sub>1],
+ Mkdir user\<^sub>2 perms\<^sub>2 [user\<^sub>1, name\<^sub>1, name\<^sub>2],
+ Creat user\<^sub>2 perms\<^sub>2 [user\<^sub>1, name\<^sub>1, name\<^sub>2, name\<^sub>3]]"
definition
- "bogus_path = [user\<^isub>1, name\<^isub>1, name\<^isub>2]"
+ "bogus_path = [user\<^sub>1, name\<^sub>1, name\<^sub>2]"
text {*
The @{term bogus} operations are the ones that lead into the uncouth
@@ -840,15 +840,15 @@
The following invariant over the root file-system describes the
bogus situation in an abstract manner: located at a certain @{term
path} within the file-system is a non-empty directory that is
- neither owned nor writable by @{term user\<^isub>1}.
+ neither owned nor writable by @{term user\<^sub>1}.
*}
definition
"invariant root path =
(\<exists>att dir.
- access root path user\<^isub>1 {} = Some (Env att dir) \<and> dir \<noteq> empty \<and>
- user\<^isub>1 \<noteq> owner att \<and>
- access root path user\<^isub>1 {Writable} = None)"
+ access root path user\<^sub>1 {} = Some (Env att dir) \<and> dir \<noteq> empty \<and>
+ user\<^sub>1 \<noteq> owner att \<and>
+ access root path user\<^sub>1 {Writable} = None)"
text {*
Following the general procedure (see
@@ -858,15 +858,15 @@
\begin{enumerate}
\item The invariant is sufficiently strong to entail the
- pathological case that @{term user\<^isub>1} is unable to remove the
- (owned) directory at @{term "[user\<^isub>1, name\<^isub>1]"}.
+ pathological case that @{term user\<^sub>1} is unable to remove the
+ (owned) directory at @{term "[user\<^sub>1, name\<^sub>1]"}.
\item The invariant does hold after having executed the @{term
bogus} list of operations (starting with an initial file-system
configuration).
\item The invariant is preserved by any file-system operation
- performed by @{term user\<^isub>1} alone, without any help by other
+ performed by @{term user\<^sub>1} alone, without any help by other
users.
\end{enumerate}
@@ -884,16 +884,16 @@
lemma cannot_rmdir:
assumes inv: "invariant root bogus_path"
- and rmdir: "root \<midarrow>(Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1])\<rightarrow> root'"
+ and rmdir: "root \<midarrow>(Rmdir user\<^sub>1 [user\<^sub>1, name\<^sub>1])\<rightarrow> root'"
shows False
proof -
- from inv obtain "file" where "access root bogus_path user\<^isub>1 {} = Some file"
+ from inv obtain "file" where "access root bogus_path user\<^sub>1 {} = Some file"
unfolding invariant_def by blast
moreover
from rmdir obtain att where
- "access root [user\<^isub>1, name\<^isub>1] user\<^isub>1 {} = Some (Env att empty)"
+ "access root [user\<^sub>1, name\<^sub>1] user\<^sub>1 {} = Some (Env att empty)"
by cases auto
- then have "access root ([user\<^isub>1, name\<^isub>1] @ [name\<^isub>2]) user\<^isub>1 {} = empty name\<^isub>2"
+ then have "access root ([user\<^sub>1, name\<^sub>1] @ [name\<^sub>2]) user\<^sub>1 {} = empty name\<^sub>2"
by (simp only: access_empty_lookup lookup_append_some) simp
ultimately show False by (simp add: bogus_path_def)
qed
@@ -916,7 +916,7 @@
text {*
\medskip At last we are left with the main effort to show that the
``bogosity'' invariant is preserved by any file-system operation
- performed by @{term user\<^isub>1} alone. Note that this holds for
+ performed by @{term user\<^sub>1} alone. Note that this holds for
any @{term path} given, the particular @{term bogus_path} is not
required here.
*}
@@ -924,20 +924,20 @@
lemma preserve_invariant:
assumes tr: "root \<midarrow>x\<rightarrow> root'"
and inv: "invariant root path"
- and uid: "uid_of x = user\<^isub>1"
+ and uid: "uid_of x = user\<^sub>1"
shows "invariant root' path"
proof -
from inv obtain att dir where
- inv1: "access root path user\<^isub>1 {} = Some (Env att dir)" and
+ inv1: "access root path user\<^sub>1 {} = Some (Env att dir)" and
inv2: "dir \<noteq> empty" and
- inv3: "user\<^isub>1 \<noteq> owner att" and
- inv4: "access root path user\<^isub>1 {Writable} = None"
+ inv3: "user\<^sub>1 \<noteq> owner att" and
+ inv4: "access root path user\<^sub>1 {Writable} = None"
by (auto simp add: invariant_def)
from inv1 have lookup: "lookup root path = Some (Env att dir)"
by (simp only: access_empty_lookup)
- from inv1 inv3 inv4 and user\<^isub>1_not_root
+ from inv1 inv3 inv4 and user\<^sub>1_not_root
have not_writable: "Writable \<notin> others att"
by (auto simp add: access_def split: option.splits)
@@ -953,7 +953,7 @@
proof (rule prefixeq_cases)
assume "path_of x \<parallel> path"
with inv root'
- have "\<And>perms. access root' path user\<^isub>1 perms = access root path user\<^isub>1 perms"
+ have "\<And>perms. access root' path user\<^sub>1 perms = access root path user\<^sub>1 perms"
by (simp only: access_update_other)
with inv show "invariant root' path"
by (simp only: invariant_def)
@@ -964,7 +964,7 @@
show ?thesis
proof (cases ys)
assume "ys = []"
- with tr uid inv2 inv3 lookup changed path and user\<^isub>1_not_root
+ with tr uid inv2 inv3 lookup changed path and user\<^sub>1_not_root
have False
by cases (auto simp add: access_empty_lookup dest: access_some_lookup)
then show ?thesis ..
@@ -1044,11 +1044,11 @@
ultimately show ?thesis ..
qed
- from lookup' have inv1': "access root' path user\<^isub>1 {} = Some (Env att dir')"
+ from lookup' have inv1': "access root' path user\<^sub>1 {} = Some (Env att dir')"
by (simp only: access_empty_lookup)
- from inv3 lookup' and not_writable user\<^isub>1_not_root
- have "access root' path user\<^isub>1 {Writable} = None"
+ from inv3 lookup' and not_writable user\<^sub>1_not_root
+ have "access root' path user\<^sub>1 {Writable} = None"
by (simp add: access_def)
with inv1' inv2' inv3 show ?thesis unfolding invariant_def by blast
qed
@@ -1066,8 +1066,8 @@
corollary
assumes bogus: "init users =bogus\<Rightarrow> root"
- shows "\<not> (\<exists>xs. (\<forall>x \<in> set xs. uid_of x = user\<^isub>1) \<and>
- can_exec root (xs @ [Rmdir user\<^isub>1 [user\<^isub>1, name\<^isub>1]]))"
+ shows "\<not> (\<exists>xs. (\<forall>x \<in> set xs. uid_of x = user\<^sub>1) \<and>
+ can_exec root (xs @ [Rmdir user\<^sub>1 [user\<^sub>1, name\<^sub>1]]))"
proof -
from cannot_rmdir init_invariant preserve_invariant
and bogus show ?thesis by (rule general_procedure)
--- a/src/HOL/ex/FinFunPred.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/ex/FinFunPred.thy Tue Aug 13 16:25:47 2013 +0200
@@ -8,7 +8,7 @@
text {* Instantiate FinFun predicates just like predicates *}
-type_synonym 'a pred\<^isub>f = "'a \<Rightarrow>f bool"
+type_synonym 'a pred\<^sub>f = "'a \<Rightarrow>f bool"
instantiation "finfun" :: (type, ord) ord
begin
@@ -101,13 +101,13 @@
Replicate predicate operations for FinFuns
*}
-abbreviation finfun_empty :: "'a pred\<^isub>f" ("{}\<^isub>f")
-where "{}\<^isub>f \<equiv> bot"
+abbreviation finfun_empty :: "'a pred\<^sub>f" ("{}\<^sub>f")
+where "{}\<^sub>f \<equiv> bot"
-abbreviation finfun_UNIV :: "'a pred\<^isub>f"
+abbreviation finfun_UNIV :: "'a pred\<^sub>f"
where "finfun_UNIV \<equiv> top"
-definition finfun_single :: "'a \<Rightarrow> 'a pred\<^isub>f"
+definition finfun_single :: "'a \<Rightarrow> 'a pred\<^sub>f"
where [code]: "finfun_single x = finfun_empty(x $:= True)"
lemma finfun_single_apply [simp]:
@@ -129,7 +129,7 @@
Warning: @{text "finfun_Ball"} and @{text "finfun_Ex"} may raise an exception, they should not be used for quickcheck
*}
-definition finfun_Ball_except :: "'a list \<Rightarrow> 'a pred\<^isub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+definition finfun_Ball_except :: "'a list \<Rightarrow> 'a pred\<^sub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
where [code del]: "finfun_Ball_except xs A P = (\<forall>a. A $ a \<longrightarrow> a \<in> set xs \<or> P a)"
lemma finfun_Ball_except_const:
@@ -149,14 +149,14 @@
shows "finfun_Ball_except xs (finfun_update_code f a b) P = ((a \<in> set xs \<or> (b \<longrightarrow> P a)) \<and> finfun_Ball_except (a # xs) f P)"
by(simp add: finfun_Ball_except_update)
-definition finfun_Ball :: "'a pred\<^isub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+definition finfun_Ball :: "'a pred\<^sub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
where [code del]: "finfun_Ball A P = Ball {x. A $ x} P"
lemma finfun_Ball_code [code]: "finfun_Ball = finfun_Ball_except []"
by(auto intro!: ext simp add: finfun_Ball_except_def finfun_Ball_def)
-definition finfun_Bex_except :: "'a list \<Rightarrow> 'a pred\<^isub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+definition finfun_Bex_except :: "'a list \<Rightarrow> 'a pred\<^sub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
where [code del]: "finfun_Bex_except xs A P = (\<exists>a. A $ a \<and> a \<notin> set xs \<and> P a)"
lemma finfun_Bex_except_const:
@@ -176,7 +176,7 @@
shows "finfun_Bex_except xs (finfun_update_code f a b) P \<longleftrightarrow> ((a \<notin> set xs \<and> b \<and> P a) \<or> finfun_Bex_except (a # xs) f P)"
by(simp add: finfun_Bex_except_update)
-definition finfun_Bex :: "'a pred\<^isub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+definition finfun_Bex :: "'a pred\<^sub>f \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
where [code del]: "finfun_Bex A P = Bex {x. A $ x} P"
lemma finfun_Bex_code [code]: "finfun_Bex = finfun_Bex_except []"
@@ -214,7 +214,7 @@
by(simp add: inf_fun_def)
lemma iso_finfun_empty_conv [code_unfold]:
- "(\<lambda>_. False) = op $ {}\<^isub>f"
+ "(\<lambda>_. False) = op $ {}\<^sub>f"
by simp
lemma iso_finfun_UNIV_conv [code_unfold]:
@@ -222,17 +222,17 @@
by simp
lemma iso_finfun_upd [code_unfold]:
- fixes A :: "'a pred\<^isub>f"
+ fixes A :: "'a pred\<^sub>f"
shows "(op $ A)(x := b) = op $ (A(x $:= b))"
by(simp add: fun_eq_iff)
lemma iso_finfun_uminus [code_unfold]:
- fixes A :: "'a pred\<^isub>f"
+ fixes A :: "'a pred\<^sub>f"
shows "- op $ A = op $ (- A)"
by(simp)
lemma iso_finfun_minus [code_unfold]:
- fixes A :: "'a pred\<^isub>f"
+ fixes A :: "'a pred\<^sub>f"
shows "op $ A - op $ B = op $ (A - B)"
by(simp)
--- a/src/HOL/ex/Groebner_Examples.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/ex/Groebner_Examples.thy Tue Aug 13 16:25:47 2013 +0200
@@ -19,7 +19,7 @@
lemma
fixes x :: int
- shows "(x - (-2))^5 = x ^ 5 + (10 * x ^ 4 + (40 * x ^ 3 + (80 * x\<twosuperior> + (80 * x + 32))))"
+ shows "(x - (-2))^5 = x ^ 5 + (10 * x ^ 4 + (40 * x ^ 3 + (80 * x\<^sup>2 + (80 * x + 32))))"
apply (tactic {* ALLGOALS (CONVERSION
(Conv.arg_conv (Conv.arg1_conv (Semiring_Normalizer.semiring_normalize_conv @{context})))) *})
by (rule refl)
@@ -54,7 +54,7 @@
lemma "(x::int)^3 - x^2 - 5*x - 3 = 0 \<longleftrightarrow> (x = 3 \<or> x = -1)"
by algebra
-theorem "x* (x\<twosuperior> - x - 5) - 3 = (0::int) \<longleftrightarrow> (x = 3 \<or> x = -1)"
+theorem "x* (x\<^sup>2 - x - 5) - 3 = (0::int) \<longleftrightarrow> (x = 3 \<or> x = -1)"
by algebra
lemma
@@ -103,7 +103,7 @@
((Ax - Bx) * (By - Cy) = (Ay - By) * (Bx - Cx))"
lemma collinear_inv_rotation:
- assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<twosuperior> + s\<twosuperior> = 1"
+ assumes "collinear (Ax, Ay) (Bx, By) (Cx, Cy)" and "c\<^sup>2 + s\<^sup>2 = 1"
shows "collinear (Ax * c - Ay * s, Ay * c + Ax * s)
(Bx * c - By * s, By * c + Bx * s) (Cx * c - Cy * s, Cy * c + Cx * s)"
using assms
--- a/src/HOL/ex/Sqrt.thy Tue Aug 13 14:20:22 2013 +0200
+++ b/src/HOL/ex/Sqrt.thy Tue Aug 13 16:25:47 2013 +0200
@@ -19,23 +19,23 @@
then obtain m n :: nat where
n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
and gcd: "gcd m n = 1" by (rule Rats_abs_nat_div_natE)
- have eq: "m\<twosuperior> = p * n\<twosuperior>"
+ have eq: "m\<^sup>2 = p * n\<^sup>2"
proof -
from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
- then have "m\<twosuperior> = (sqrt p)\<twosuperior> * n\<twosuperior>"
+ then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2"
by (auto simp add: power2_eq_square)
- also have "(sqrt p)\<twosuperior> = p" by simp
- also have "\<dots> * n\<twosuperior> = p * n\<twosuperior>" by simp
+ also have "(sqrt p)\<^sup>2 = p" by simp
+ also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
finally show ?thesis ..
qed
have "p dvd m \<and> p dvd n"
proof
- from eq have "p dvd m\<twosuperior>" ..
+ from eq have "p dvd m\<^sup>2" ..
with `prime p` pos2 show "p dvd m" by (rule prime_dvd_power_nat)
then obtain k where "m = p * k" ..
- with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
- with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
- then have "p dvd n\<twosuperior>" ..
+ with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square mult_ac)
+ with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
+ then have "p dvd n\<^sup>2" ..
with `prime p` pos2 show "p dvd n" by (rule prime_dvd_power_nat)
qed
then have "p dvd gcd m n" ..
@@ -65,17 +65,17 @@
n: "n \<noteq> 0" and sqrt_rat: "\<bar>sqrt p\<bar> = m / n"
and gcd: "gcd m n = 1" by (rule Rats_abs_nat_div_natE)
from n and sqrt_rat have "m = \<bar>sqrt p\<bar> * n" by simp
- then have "m\<twosuperior> = (sqrt p)\<twosuperior> * n\<twosuperior>"
+ then have "m\<^sup>2 = (sqrt p)\<^sup>2 * n\<^sup>2"
by (auto simp add: power2_eq_square)
- also have "(sqrt p)\<twosuperior> = p" by simp
- also have "\<dots> * n\<twosuperior> = p * n\<twosuperior>" by simp
- finally have eq: "m\<twosuperior> = p * n\<twosuperior>" ..
- then have "p dvd m\<twosuperior>" ..
+ also have "(sqrt p)\<^sup>2 = p" by simp
+ also have "\<dots> * n\<^sup>2 = p * n\<^sup>2" by simp
+ finally have eq: "m\<^sup>2 = p * n\<^sup>2" ..
+ then have "p dvd m\<^sup>2" ..
with `prime p` pos2 have dvd_m: "p dvd m" by (rule prime_dvd_power_nat)
then obtain k where "m = p * k" ..
- with eq have "p * n\<twosuperior> = p\<twosuperior> * k\<twosuperior>" by (auto simp add: power2_eq_square mult_ac)
- with p have "n\<twosuperior> = p * k\<twosuperior>" by (simp add: power2_eq_square)
- then have "p dvd n\<twosuperior>" ..
+ with eq have "p * n\<^sup>2 = p\<^sup>2 * k\<^sup>2" by (auto simp add: power2_eq_square mult_ac)
+ with p have "n\<^sup>2 = p * k\<^sup>2" by (simp add: power2_eq_square)
+ then have "p dvd n\<^sup>2" ..
with `prime p` pos2 have "p dvd n" by (rule prime_dvd_power_nat)
with dvd_m have "p dvd gcd m n" by (rule gcd_greatest_nat)
with gcd have "p dvd 1" by simp